Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 16d40d7

Browse files
committed
feat(data/finset): fintype, multiset.sort, list.pmap
1 parent c03c16d commit 16d40d7

File tree

15 files changed

+558
-224
lines changed

15 files changed

+558
-224
lines changed

algebra/big_operators.lean

Lines changed: 10 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -61,10 +61,11 @@ lemma prod_sdiff (h : s₁ ⊆ s₂) : (s₂ \ s₁).prod f * s₁.prod f = s₂
6161
by rw [←prod_union (sdiff_inter_self _ _), sdiff_union_of_subset h]
6262

6363
@[to_additive finset.sum_bind]
64-
lemma prod_bind [decidable_eq γ] {s : finset γ} {t : γ → finset α} :
64+
lemma prod_bind {s : finset γ} {t : γ → finset α} :
6565
(∀x∈s, ∀y∈s, x ≠ y → t x ∩ t y = ∅) → (s.bind t).prod f = s.prod (λx, (t x).prod f) :=
66-
finset.induction_on s (by simp) $
67-
assume x s hxs ih hd,
66+
by have := classical.decidable_eq γ; exact
67+
finset.induction_on s (by simp)
68+
(assume x s hxs ih hd,
6869
have hd' : ∀x∈s, ∀y∈s, x ≠ y → t x ∩ t y = ∅,
6970
from assume _ hx _ hy, hd _ (mem_insert_of_mem hx) _ (mem_insert_of_mem hy),
7071
have t x ∩ finset.bind s t = ∅,
@@ -75,12 +76,13 @@ finset.induction_on s (by simp) $
7576
have t x ∩ t y = ∅,
7677
from hd _ (mem_insert_self _ _) _ (mem_insert_of_mem hys) $ assume h, hxs $ h.symm ▸ hys,
7778
by rwa [this] at ha,
78-
by simp [hxs, prod_union this, ih hd'] {contextual := tt}
79+
by simp [hxs, prod_union this, ih hd'] {contextual := tt})
7980

8081
@[to_additive finset.sum_product]
81-
lemma prod_product [decidable_eq γ] {s : finset γ} {t : finset α} {f : γ×α → β} :
82+
lemma prod_product {s : finset γ} {t : finset α} {f : γ×α → β} :
8283
(s.product t).prod f = (s.prod $ λx, t.prod $ λy, f (x, y)) :=
8384
begin
85+
have := classical.decidable_eq γ,
8486
rw [product_eq_bind, prod_bind (λ x hx y hy h, ext.2 _)], {simp [prod_image]},
8587
simp [mem_image], intros, intro, refine h _, cc
8688
end
@@ -93,8 +95,9 @@ have ∀a₁ a₂:α, ∀s₁ : finset (σ a₁), ∀s₂ : finset (σ a₂), a
9395
s₁.image (sigma.mk a₁) ∩ s₂.image (sigma.mk a₂) = ∅,
9496
from assume b₁ b₂ s₁ s₂ h, ext.2 $ assume ⟨b₃, c₃⟩,
9597
by simp [mem_image, sigma.mk_eq_mk_iff, heq_iff_eq] {contextual := tt}; cc,
96-
calc (s.bind (λa, (t a).image (λs, sigma.mk a s))).prod f =
97-
s.prod (λa, ((t a).image (λs, sigma.mk a s)).prod f) :
98+
calc (s.sigma t).prod f =
99+
(s.bind (λa, (t a).image (λs, sigma.mk a s))).prod f : by rw sigma_eq_bind
100+
... = s.prod (λa, ((t a).image (λs, sigma.mk a s)).prod f) :
98101
prod_bind $ assume a₁ ha a₂ ha₂ h, this a₁ a₂ (t a₁) (t a₂) h
99102
... = (s.prod $ λa, (t a).prod $ λs, f ⟨a, s⟩) :
100103
by simp [prod_image, sigma.mk_eq_mk_iff, heq_iff_eq]

analysis/measure_theory/outer_measure.lean

Lines changed: 3 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -12,11 +12,10 @@ import data.set order.galois_connection algebra.big_operators
1212

1313
noncomputable theory
1414

15-
open classical set lattice finset function filter
15+
open set lattice finset function filter
1616
open ennreal (of_real)
17-
local attribute [instance] decidable_inhabited prop_decidable
17+
local attribute [instance] classical.prop_decidable
1818

19-
open classical
2019
local infix ` ^ ` := monoid.pow
2120

2221
namespace measure_theory
@@ -95,7 +94,7 @@ let μ := λs, ⨅{f : ℕ → set α} (h : s ⊆ ⋃i, f i), ∑i, m (f i) in
9594
... < ⊤ : hb)
9695
(by simp; exact hε' _),
9796
by simpa [μ, infi_lt_iff] using this,
98-
let ⟨f, hf⟩ := axiom_of_choice this in
97+
let ⟨f, hf⟩ := classical.axiom_of_choice this in
9998
let f' := λi, f (nat.unpair i).1 (nat.unpair i).2 in
10099
have hf' : (⋃ (i : ℕ), s i) ⊆ (⋃i, f' i),
101100
from Union_subset $ assume i, subset.trans (hf i).left $ Union_subset_Union2 $ assume j,

analysis/topology/continuity.lean

Lines changed: 4 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -13,8 +13,8 @@ A major difference is that this formalization is heavily based on the filter lib
1313
import analysis.topology.topological_space
1414
noncomputable theory
1515

16-
open set filter lattice classical
17-
local attribute [instance] decidable_inhabited prop_decidable
16+
open set filter lattice
17+
local attribute [instance] classical.prop_decidable
1818

1919
universes u v w x y
2020
variables {α : Type u} {β : Type v} {γ : Type w} {δ : Type y} {ι : Sort x}
@@ -61,14 +61,14 @@ assume s hs,
6161
have (λa, ite (p a) (f a) (g a)) ⁻¹' s =
6262
(closure {a | p a} ∩ f ⁻¹' s) ∪ (closure {a | ¬ p a} ∩ g ⁻¹' s),
6363
from set.ext $ assume a,
64-
by_cases
64+
classical.by_cases
6565
(assume : a ∈ frontier {a | p a},
6666
have hac : a ∈ closure {a | p a}, from this.left,
6767
have hai : a ∈ closure {a | ¬ p a},
6868
from have a ∈ - interior {a | p a}, from this.right, by rwa [←closure_compl] at this,
6969
by by_cases p a; simp [h, hp a this, hac, hai, iff_def] {contextual := tt})
7070
(assume hf : a ∈ - frontier {a | p a},
71-
by_cases
71+
classical.by_cases
7272
(assume : p a,
7373
have hc : a ∈ closure {a | p a}, from subset_closure this,
7474
have hnc : a ∉ closure {a | ¬ p a},

data/encodable.lean

Lines changed: 2 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -197,11 +197,10 @@ encodable_of_equiv {s : multiset α // s.nodup}
197197
⟨λ ⟨a, b⟩, ⟨a, b⟩, λ⟨a, b⟩, ⟨a, b⟩, λ ⟨a, b⟩, rfl, λ⟨a, b⟩, rfl⟩
198198

199199
instance encodable_ulift [encodable α] : encodable (ulift α) :=
200-
encodable_of_equiv α ⟨ulift.down, ulift.up, ulift.up_down, ulift.down_up⟩
200+
encodable_of_equiv _ equiv.ulift
201201

202202
instance encodable_plift [encodable α] : encodable (plift α) :=
203-
encodable_of_equiv α ⟨plift.down, plift.up, plift.up_down, plift.down_up⟩
204-
203+
encodable_of_equiv _ equiv.plift
205204

206205
noncomputable def encodable_of_inj [encodable β] (f : α → β) (hf : injective f) : encodable α :=
207206
encodable_of_left_injection f (partial_inv f) (partial_inv_eq hf)

data/equiv.lean

Lines changed: 43 additions & 32 deletions
Original file line numberDiff line numberDiff line change
@@ -8,7 +8,7 @@ We say two types are equivalent if they are isomorphic.
88
99
Two equivalent types have the same cardinality.
1010
-/
11-
import data.prod data.nat.pairing tactic
11+
import data.prod data.nat.pairing logic.function tactic
1212
open function
1313

1414
universes u v w
@@ -20,22 +20,22 @@ def map {p : α → Prop} {q : β → Prop} (f : α → β) (h : ∀a, p a → q
2020
subtype p → subtype q
2121
| ⟨v, hv⟩ := ⟨f v, h v hv⟩
2222

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}
2424
(f : α → β) (h : ∀a, p a → q (f a)) (g : β → γ) (l : ∀a, q a → r (g a)) :
2525
map g l (map f h x) = map (g ∘ f) (assume a ha, l (f a) $ h a ha) x :=
2626
by cases x with v h; refl
2727

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 :=
2929
funext $ assume ⟨v, h⟩, rfl
3030

3131
end subtype
3232

3333
namespace function
3434

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 :=
3636
funext $ h
3737

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 :=
3939
funext $ h
4040

4141
end function
@@ -54,10 +54,10 @@ infix ` ≃ `:50 := equiv
5454
instance : has_coe_to_fun (α ≃ β) :=
5555
⟨_, to_fun⟩
5656

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 :=
5858
rfl
5959

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₂
6161
| ⟨f₁, g₁, l₁, r₁⟩ ⟨f₂, g₂, l₂, r₂⟩ h :=
6262
have f₁ = f₂, from h,
6363
have g₁ = g₂, from funext $ assume x,
@@ -78,30 +78,31 @@ lemma eq_of_to_fun_eq : ∀ {e₁ e₂ : equiv α β}, (e₁ : α → β) = e₂
7878
λ x, show g₁ (g₂ (f₂ (f₁ x))) = x, by rw [l₂, l₁],
7979
λ x, show f₂ (f₁ (g₁ (g₂ x))) = x, by rw [r₁, r₂]⟩
8080

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+
8185
def id := equiv.refl α
8286

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 :=
8488
rfl
8589

86-
lemma id_apply (x : α) : @id α x = x :=
90+
theorem id_apply (x : α) : @id α x = x :=
8791
rfl
8892

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) :=
9094
by cases g; cases f; simp [equiv.trans, *]
9195

92-
lemma apply_inverse_apply : ∀ (e : α ≃ β) (x : β), e (e.symm x) = x
96+
theorem apply_inverse_apply : ∀ (e : α ≃ β) (x : β), e (e.symm x) = x
9397
| ⟨f₁, g₁, l₁, r₁⟩ x := by simp [equiv.symm]; rw r₁
9498

95-
lemma inverse_apply_apply : ∀ (e : α ≃ β) (x : α), e.symm (e x) = x
99+
theorem inverse_apply_apply : ∀ (e : α ≃ β) (x : α), e.symm (e x) = x
96100
| ⟨f₁, g₁, l₁, r₁⟩ x := by simp [equiv.symm]; rw l₁
97101

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
103104

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
105106
| ⟨f₁, g₁, l₁, r₁⟩ x y := by simp [equiv.symm];
106107
show f₁ x = y ↔ x = g₁ y; from
107108
⟨λ e : f₁ x = y, e ▸ (l₁ x).symm,
@@ -113,12 +114,15 @@ def equiv_empty (h : α → false) : α ≃ empty :=
113114
def false_equiv_empty : false ≃ empty :=
114115
equiv_empty _root_.id
115116

116-
lemma empty_of_not_nonempty {α : Sort*} (h : ¬ nonempty α) : α ≃ empty :=
117+
def empty_of_not_nonempty {α : Sort*} (h : ¬ nonempty α) : α ≃ empty :=
117118
⟨assume a, (h ⟨a⟩).elim, assume e, e.rec_on _, assume a, (h ⟨a⟩).elim, assume e, e.rec_on _⟩
118119

119-
protected lemma ulift {α : Type u} : ulift α ≃ α :=
120+
protected def ulift {α : Type u} : ulift α ≃ α :=
120121
⟨ulift.down, ulift.up, ulift.up_down, ulift.down_up⟩
121122

123+
protected def plift {α : Type u} : plift α ≃ α :=
124+
⟨plift.down, plift.up, plift.up_down, plift.down_up⟩
125+
122126
@[congr] def arrow_congr {α₁ β₁ α₂ β₂ : Sort*} : α₁ ≃ α₂ → β₁ ≃ β₂ → (α₁ → β₁) ≃ (α₂ → β₂)
123127
| ⟨f₁, g₁, l₁, r₁⟩ ⟨f₂, g₂, l₂, r₂⟩ :=
124128
⟨λ (h : α₁ → β₁) (a : α₂), f₂ (h (g₁ a)),
@@ -140,7 +144,7 @@ section
140144
calc (false → α) ≃ (empty → α) : arrow_congr false_equiv_empty (equiv.refl _)
141145
... ≃ unit : empty_arrow_equiv_unit _
142146

143-
lemma arrow_empty_unit {α : Sort*} : (empty → α) ≃ unit :=
147+
def arrow_empty_unit {α : Sort*} : (empty → α) ≃ unit :=
144148
⟨λf, (), λu e, e.rec_on _, assume f, funext $ assume e, e.rec_on _, assume u, unit_eq _ _⟩
145149

146150
end
@@ -219,6 +223,13 @@ def bool_equiv_unit_sum_unit : bool ≃ (unit ⊕ unit) :=
219223
λ s, match s with inr _ := none | inl a := some a end,
220224
λ o, by cases o; refl,
221225
λ 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+
222233
end
223234

224235
section
@@ -405,10 +416,10 @@ if r = a then b
405416
else if r = b then a
406417
else r
407418

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 :=
409420
by by_cases r = a; simp [swap_core, *]
410421

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 :=
412423
begin
413424
by_cases r = b with hb,
414425
{ by_cases r = a with ha,
@@ -421,7 +432,7 @@ begin
421432
simp [swap_core, *] }
422433
end
423434

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 :=
425436
begin
426437
by_cases r = b with hb,
427438
{ by_cases r = a with ha,
@@ -437,28 +448,28 @@ end
437448
def swap (a b : α) : perm α :=
438449
⟨swap_core a b, swap_core a b, λr, swap_core_swap_core r a b, λr, swap_core_swap_core r a b⟩
439450

440-
lemma swap_self (a : α) : swap a a = id :=
451+
theorem swap_self (a : α) : swap a a = id :=
441452
eq_of_to_fun_eq $ funext $ λ r, swap_core_self r a
442453

443-
lemma swap_comm (a b : α) : swap a b = swap b a :=
454+
theorem swap_comm (a b : α) : swap a b = swap b a :=
444455
eq_of_to_fun_eq $ funext $ λ r, swap_core_comm r _ _
445456

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 :=
447458
rfl
448459

449-
lemma swap_apply_left (a b : α) : swap a b a = b :=
460+
theorem swap_apply_left (a b : α) : swap a b a = b :=
450461
if_pos rfl
451462

452-
lemma swap_apply_right (a b : α) : swap a b b = a :=
463+
theorem swap_apply_right (a b : α) : swap a b b = a :=
453464
by by_cases b = a; simp [swap_apply_def, *]
454465

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 :=
456467
by simp [swap_apply_def] {contextual := tt}
457468

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 :=
459470
eq_of_to_fun_eq $ funext $ λ x, swap_core_swap_core _ _ _
460471

461-
lemma swap_comp_apply {a b x : α} (π : perm α) :
472+
theorem swap_comp_apply {a b x : α} (π : perm α) :
462473
π.trans (swap a b) x = if π x = a then b else if π x = b then a else π x :=
463474
by cases π; refl
464475

0 commit comments

Comments
 (0)