@@ -8,7 +8,7 @@ We say two types are equivalent if they are isomorphic.
8
8
9
9
Two equivalent types have the same cardinality.
10
10
-/
11
- import data.prod data.nat.pairing tactic
11
+ import data.prod data.nat.pairing logic.function tactic
12
12
open function
13
13
14
14
universes u v w
@@ -20,22 +20,22 @@ def map {p : α → Prop} {q : β → Prop} (f : α → β) (h : ∀a, p a → q
20
20
subtype p → subtype q
21
21
| ⟨v, hv⟩ := ⟨f v, h v hv⟩
22
22
23
- lemma map_comp {p : α → Prop } {q : β → Prop } {r : γ → Prop } {x : subtype p}
23
+ theorem map_comp {p : α → Prop } {q : β → Prop } {r : γ → Prop } {x : subtype p}
24
24
(f : α → β) (h : ∀a, p a → q (f a)) (g : β → γ) (l : ∀a, q a → r (g a)) :
25
25
map g l (map f h x) = map (g ∘ f) (assume a ha, l (f a) $ h a ha) x :=
26
26
by cases x with v h; refl
27
27
28
- lemma map_id {p : α → Prop } {h : ∀a, p a → p (id a)} : map (@id α) h = id :=
28
+ theorem map_id {p : α → Prop } {h : ∀a, p a → p (id a)} : map (@id α) h = id :=
29
29
funext $ assume ⟨v, h⟩, rfl
30
30
31
31
end subtype
32
32
33
33
namespace function
34
34
35
- lemma left_inverse.f_g_eq_id {f : α → β} {g : β → α} (h : left_inverse f g) : f ∘ g = id :=
35
+ theorem left_inverse.f_g_eq_id {f : α → β} {g : β → α} (h : left_inverse f g) : f ∘ g = id :=
36
36
funext $ h
37
37
38
- lemma right_inverse.g_f_eq_id {f : α → β} {g : β → α} (h : right_inverse f g) : g ∘ f = id :=
38
+ theorem right_inverse.g_f_eq_id {f : α → β} {g : β → α} (h : right_inverse f g) : g ∘ f = id :=
39
39
funext $ h
40
40
41
41
end function
@@ -54,10 +54,10 @@ infix ` ≃ `:50 := equiv
54
54
instance : has_coe_to_fun (α ≃ β) :=
55
55
⟨_, to_fun⟩
56
56
57
- @[simp] lemma coe_fn_mk (f : α → β) (g l r) : (equiv.mk f g l r : α → β) = f :=
57
+ @[simp] theorem coe_fn_mk (f : α → β) (g l r) : (equiv.mk f g l r : α → β) = f :=
58
58
rfl
59
59
60
- lemma eq_of_to_fun_eq : ∀ {e₁ e₂ : equiv α β}, (e₁ : α → β) = e₂ → e₁ = e₂
60
+ theorem eq_of_to_fun_eq : ∀ {e₁ e₂ : equiv α β}, (e₁ : α → β) = e₂ → e₁ = e₂
61
61
| ⟨f₁, g₁, l₁, r₁⟩ ⟨f₂, g₂, l₂, r₂⟩ h :=
62
62
have f₁ = f₂, from h,
63
63
have g₁ = g₂, from funext $ assume x,
@@ -78,30 +78,31 @@ lemma eq_of_to_fun_eq : ∀ {e₁ e₂ : equiv α β}, (e₁ : α → β) = e₂
78
78
λ x, show g₁ (g₂ (f₂ (f₁ x))) = x, by rw [l₂, l₁],
79
79
λ x, show f₂ (f₁ (g₁ (g₂ x))) = x, by rw [r₁, r₂]⟩
80
80
81
+ protected theorem bijective : ∀ f : α ≃ β, bijective f
82
+ | ⟨f, g, h₁, h₂⟩ :=
83
+ ⟨injective_of_left_inverse h₁, surjective_of_has_right_inverse ⟨_, h₂⟩⟩
84
+
81
85
def id := equiv.refl α
82
86
83
- @[simp] lemma coe_fn_symm_mk (f : α → β) (g l r) : ((equiv.mk f g l r).symm : β → α) = g :=
87
+ @[simp] theorem coe_fn_symm_mk (f : α → β) (g l r) : ((equiv.mk f g l r).symm : β → α) = g :=
84
88
rfl
85
89
86
- lemma id_apply (x : α) : @id α x = x :=
90
+ theorem id_apply (x : α) : @id α x = x :=
87
91
rfl
88
92
89
- lemma comp_apply (g : β ≃ γ) (f : α ≃ β) (x : α) : (g ∘ f) x = g (f x) :=
93
+ theorem comp_apply (g : β ≃ γ) (f : α ≃ β) (x : α) : (g ∘ f) x = g (f x) :=
90
94
by cases g; cases f; simp [equiv.trans, *]
91
95
92
- lemma apply_inverse_apply : ∀ (e : α ≃ β) (x : β), e (e.symm x) = x
96
+ theorem apply_inverse_apply : ∀ (e : α ≃ β) (x : β), e (e.symm x) = x
93
97
| ⟨f₁, g₁, l₁, r₁⟩ x := by simp [equiv.symm]; rw r₁
94
98
95
- lemma inverse_apply_apply : ∀ (e : α ≃ β) (x : α), e.symm (e x) = x
99
+ theorem inverse_apply_apply : ∀ (e : α ≃ β) (x : α), e.symm (e x) = x
96
100
| ⟨f₁, g₁, l₁, r₁⟩ x := by simp [equiv.symm]; rw l₁
97
101
98
- lemma eq_iff_eq_of_injective {f : α → β} (inj : injective f) (a b : α) : f a = f b ↔ a = b :=
99
- ⟨λ h, inj h, λ h, by rw h⟩
100
-
101
- lemma apply_eq_iff_eq : ∀ (f : α ≃ β) (x y : α), f x = f y ↔ x = y
102
- | ⟨f₁, g₁, l₁, r₁⟩ x y := eq_iff_eq_of_injective (injective_of_left_inverse l₁) x y
102
+ theorem apply_eq_iff_eq : ∀ (f : α ≃ β) (x y : α), f x = f y ↔ x = y
103
+ | ⟨f₁, g₁, l₁, r₁⟩ x y := (injective_of_left_inverse l₁).eq_iff
103
104
104
- lemma apply_eq_iff_eq_inverse_apply : ∀ (f : α ≃ β) (x : α) (y : β), f x = y ↔ x = f.symm y
105
+ theorem apply_eq_iff_eq_inverse_apply : ∀ (f : α ≃ β) (x : α) (y : β), f x = y ↔ x = f.symm y
105
106
| ⟨f₁, g₁, l₁, r₁⟩ x y := by simp [equiv.symm];
106
107
show f₁ x = y ↔ x = g₁ y; from
107
108
⟨λ e : f₁ x = y, e ▸ (l₁ x).symm,
@@ -113,12 +114,15 @@ def equiv_empty (h : α → false) : α ≃ empty :=
113
114
def false_equiv_empty : false ≃ empty :=
114
115
equiv_empty _root_.id
115
116
116
- lemma empty_of_not_nonempty {α : Sort *} (h : ¬ nonempty α) : α ≃ empty :=
117
+ def empty_of_not_nonempty {α : Sort *} (h : ¬ nonempty α) : α ≃ empty :=
117
118
⟨assume a, (h ⟨a⟩).elim, assume e, e.rec_on _, assume a, (h ⟨a⟩).elim, assume e, e.rec_on _⟩
118
119
119
- protected lemma ulift {α : Type u} : ulift α ≃ α :=
120
+ protected def ulift {α : Type u} : ulift α ≃ α :=
120
121
⟨ulift.down, ulift.up, ulift.up_down, ulift.down_up⟩
121
122
123
+ protected def plift {α : Type u} : plift α ≃ α :=
124
+ ⟨plift.down, plift.up, plift.up_down, plift.down_up⟩
125
+
122
126
@[congr] def arrow_congr {α₁ β₁ α₂ β₂ : Sort *} : α₁ ≃ α₂ → β₁ ≃ β₂ → (α₁ → β₁) ≃ (α₂ → β₂)
123
127
| ⟨f₁, g₁, l₁, r₁⟩ ⟨f₂, g₂, l₂, r₂⟩ :=
124
128
⟨λ (h : α₁ → β₁) (a : α₂), f₂ (h (g₁ a)),
@@ -140,7 +144,7 @@ section
140
144
calc (false → α) ≃ (empty → α) : arrow_congr false_equiv_empty (equiv.refl _)
141
145
... ≃ unit : empty_arrow_equiv_unit _
142
146
143
- lemma arrow_empty_unit {α : Sort *} : (empty → α) ≃ unit :=
147
+ def arrow_empty_unit {α : Sort *} : (empty → α) ≃ unit :=
144
148
⟨λf, (), λu e, e.rec_on _, assume f, funext $ assume e, e.rec_on _, assume u, unit_eq _ _⟩
145
149
146
150
end
@@ -219,6 +223,13 @@ def bool_equiv_unit_sum_unit : bool ≃ (unit ⊕ unit) :=
219
223
λ s, match s with inr _ := none | inl a := some a end ,
220
224
λ o, by cases o; refl,
221
225
λ s, by rcases s with _ | ⟨⟨⟩⟩; refl⟩
226
+
227
+ def sum_equiv_sigma_bool (α β : Sort *) : (α ⊕ β) ≃ (Σ b: bool, cond b α β) :=
228
+ ⟨λ s, match s with inl a := ⟨tt, a⟩ | inr b := ⟨ff, b⟩ end ,
229
+ λ s, match s with ⟨tt, a⟩ := inl a | ⟨ff, b⟩ := inr b end ,
230
+ λ s, by cases s; refl,
231
+ λ s, by rcases s with ⟨_|_, _⟩; refl⟩
232
+
222
233
end
223
234
224
235
section
@@ -405,10 +416,10 @@ if r = a then b
405
416
else if r = b then a
406
417
else r
407
418
408
- lemma swap_core_self (r a : α) : swap_core a a r = r :=
419
+ theorem swap_core_self (r a : α) : swap_core a a r = r :=
409
420
by by_cases r = a; simp [swap_core, *]
410
421
411
- lemma swap_core_swap_core (r a b : α) : swap_core a b (swap_core a b r) = r :=
422
+ theorem swap_core_swap_core (r a b : α) : swap_core a b (swap_core a b r) = r :=
412
423
begin
413
424
by_cases r = b with hb,
414
425
{ by_cases r = a with ha,
@@ -421,7 +432,7 @@ begin
421
432
simp [swap_core, *] }
422
433
end
423
434
424
- lemma swap_core_comm (r a b : α) : swap_core a b r = swap_core b a r :=
435
+ theorem swap_core_comm (r a b : α) : swap_core a b r = swap_core b a r :=
425
436
begin
426
437
by_cases r = b with hb,
427
438
{ by_cases r = a with ha,
@@ -437,28 +448,28 @@ end
437
448
def swap (a b : α) : perm α :=
438
449
⟨swap_core a b, swap_core a b, λr, swap_core_swap_core r a b, λr, swap_core_swap_core r a b⟩
439
450
440
- lemma swap_self (a : α) : swap a a = id :=
451
+ theorem swap_self (a : α) : swap a a = id :=
441
452
eq_of_to_fun_eq $ funext $ λ r, swap_core_self r a
442
453
443
- lemma swap_comm (a b : α) : swap a b = swap b a :=
454
+ theorem swap_comm (a b : α) : swap a b = swap b a :=
444
455
eq_of_to_fun_eq $ funext $ λ r, swap_core_comm r _ _
445
456
446
- lemma swap_apply_def (a b x : α) : swap a b x = if x = a then b else if x = b then a else x :=
457
+ theorem swap_apply_def (a b x : α) : swap a b x = if x = a then b else if x = b then a else x :=
447
458
rfl
448
459
449
- lemma swap_apply_left (a b : α) : swap a b a = b :=
460
+ theorem swap_apply_left (a b : α) : swap a b a = b :=
450
461
if_pos rfl
451
462
452
- lemma swap_apply_right (a b : α) : swap a b b = a :=
463
+ theorem swap_apply_right (a b : α) : swap a b b = a :=
453
464
by by_cases b = a; simp [swap_apply_def, *]
454
465
455
- lemma swap_apply_of_ne_of_ne {a b x : α} : x ≠ a → x ≠ b → swap a b x = x :=
466
+ theorem swap_apply_of_ne_of_ne {a b x : α} : x ≠ a → x ≠ b → swap a b x = x :=
456
467
by simp [swap_apply_def] {contextual := tt}
457
468
458
- lemma swap_swap (a b : α) : (swap a b).trans (swap a b) = id :=
469
+ theorem swap_swap (a b : α) : (swap a b).trans (swap a b) = id :=
459
470
eq_of_to_fun_eq $ funext $ λ x, swap_core_swap_core _ _ _
460
471
461
- lemma swap_comp_apply {a b x : α} (π : perm α) :
472
+ theorem swap_comp_apply {a b x : α} (π : perm α) :
462
473
π.trans (swap a b) x = if π x = a then b else if π x = b then a else π x :=
463
474
by cases π; refl
464
475
0 commit comments