generated from pitmonticone/LeanProject
-
Notifications
You must be signed in to change notification settings - Fork 1
/
S03_Negation.lean
138 lines (99 loc) · 2.63 KB
/
S03_Negation.lean
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
import LeanInVienna.Common
import Mathlib.Data.Real.Basic
namespace C03S03
section
variable (a b : ℝ)
example (h : a < b) : ¬b < a := by
intro h'
have : a < a := lt_trans h h'
exact lt_irrefl a this
def FnUb (f : ℝ → ℝ) (a : ℝ) : Prop :=
∀ x, f x ≤ a
def FnLb (f : ℝ → ℝ) (a : ℝ) : Prop :=
∀ x, a ≤ f x
def FnHasUb (f : ℝ → ℝ) :=
∃ a, FnUb f a
def FnHasLb (f : ℝ → ℝ) :=
∃ a, FnLb f a
variable (f : ℝ → ℝ)
example (h : ∀ a, ∃ x, f x > a) : ¬FnHasUb f := by
intro fnub
rcases fnub with ⟨a, fnuba⟩
rcases h a with ⟨x, hx⟩
have : f x ≤ a := fnuba x
linarith
example (h : ∀ a, ∃ x, f x < a) : ¬FnHasLb f :=
sorry
example : ¬FnHasUb fun x ↦ x :=
sorry
#check (not_le_of_gt : a > b → ¬a ≤ b)
#check (not_lt_of_ge : a ≥ b → ¬a < b)
#check (lt_of_not_ge : ¬a ≥ b → a < b)
#check (le_of_not_gt : ¬a > b → a ≤ b)
example (h : Monotone f) (h' : f a < f b) : a < b := by
sorry
example (h : a ≤ b) (h' : f b < f a) : ¬Monotone f := by
sorry
example : ¬∀ {f : ℝ → ℝ}, Monotone f → ∀ {a b}, f a ≤ f b → a ≤ b := by
intro h
let f := fun x : ℝ ↦ (0 : ℝ)
have monof : Monotone f := by sorry
have h' : f 1 ≤ f 0 := le_refl _
sorry
example (x : ℝ) (h : ∀ ε > 0, x < ε) : x ≤ 0 := by
sorry
end
section
variable {α : Type*} (P : α → Prop) (Q : Prop)
example (h : ¬∃ x, P x) : ∀ x, ¬P x := by
sorry
example (h : ∀ x, ¬P x) : ¬∃ x, P x := by
sorry
example (h : ¬∀ x, P x) : ∃ x, ¬P x := by
sorry
example (h : ∃ x, ¬P x) : ¬∀ x, P x := by
sorry
example (h : ¬∀ x, P x) : ∃ x, ¬P x := by
by_contra h'
apply h
intro x
show P x
by_contra h''
exact h' ⟨x, h''⟩
example (h : ¬¬Q) : Q := by
sorry
example (h : Q) : ¬¬Q := by
sorry
end
section
variable (f : ℝ → ℝ)
example (h : ¬FnHasUb f) : ∀ a, ∃ x, f x > a := by
sorry
example (h : ¬∀ a, ∃ x, f x > a) : FnHasUb f := by
push_neg at h
exact h
example (h : ¬FnHasUb f) : ∀ a, ∃ x, f x > a := by
dsimp only [FnHasUb, FnUb] at h
push_neg at h
exact h
example (h : ¬Monotone f) : ∃ x y, x ≤ y ∧ f y < f x := by
sorry
example (h : ¬FnHasUb f) : ∀ a, ∃ x, f x > a := by
contrapose! h
exact h
example (x : ℝ) (h : ∀ ε > 0, x ≤ ε) : x ≤ 0 := by
contrapose! h
use x / 2
constructor <;> linarith
end
section
variable (a : ℕ)
example (h : 0 < 0) : a > 37 := by
exfalso
apply lt_irrefl 0 h
example (h : 0 < 0) : a > 37 :=
absurd h (lt_irrefl 0)
example (h : 0 < 0) : a > 37 := by
have : ¬0 < 0 := lt_irrefl 0
contradiction
end