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

Commit 14dcfe0

Browse files
committed
chore(*): assorted lemmas (#4566)
Non-bc changes: * make some lemmas use `coe` instead of `subtype.val`; * make the arguments of `range_comp` explicit, reorder them.
1 parent 918e5d8 commit 14dcfe0

File tree

12 files changed

+91
-37
lines changed

12 files changed

+91
-37
lines changed

src/data/dfinsupp.lean

Lines changed: 8 additions & 11 deletions
Original file line numberDiff line numberDiff line change
@@ -219,7 +219,7 @@ rfl
219219

220220
@[simp] lemma subtype_domain_apply [Π i, has_zero (β i)] {p : ι → Prop} [decidable_pred p]
221221
{i : subtype p} {v : Π₀ i, β i} :
222-
(subtype_domain p v) i = v (i.val) :=
222+
(subtype_domain p v) i = v i :=
223223
quotient.induction_on v $ λ x, rfl
224224

225225
@[simp] lemma subtype_domain_add [Π i, add_monoid (β i)] {p : ι → Prop} [decidable_pred p]
@@ -580,15 +580,13 @@ by ext i; by_cases h1 : p i; by_cases h2 : f i ≠ 0;
580580
by ext i; by_cases h : p i; simp [h]
581581

582582
lemma subtype_domain_def (f : Π₀ i, β i) :
583-
f.subtype_domain p = mk (f.support.subtype p) (λ i, f i.1) :=
584-
by ext i; cases i with i hi;
585-
by_cases h1 : p i; by_cases h2 : f i ≠ 0;
586-
try {simp at h2}; dsimp; simp [h1, h2]
583+
f.subtype_domain p = mk (f.support.subtype p) (λ i, f i) :=
584+
by ext i; by_cases h1 : p i; by_cases h2 : f i ≠ 0;
585+
try {simp at h2}; dsimp; simp [h1, h2, ← subtype.val_eq_coe]
587586

588587
@[simp] lemma support_subtype_domain {f : Π₀ i, β i} :
589588
(subtype_domain p f).support = f.support.subtype p :=
590-
by ext i; cases i with i hi;
591-
by_cases h1 : p i; by_cases h2 : f i ≠ 0;
589+
by ext i; by_cases h1 : p i; by_cases h2 : f i ≠ 0;
592590
try {simp at h2}; dsimp; simp [h1, h2]
593591

594592
end filter_and_subtype_domain
@@ -790,10 +788,9 @@ end
790788
lemma prod_subtype_domain_index [Π i, has_zero (β i)] [Π i (x : β i), decidable (x ≠ 0)]
791789
[comm_monoid γ] {v : Π₀ i, β i} {p : ι → Prop} [decidable_pred p]
792790
{h : Π i, β i → γ} (hp : ∀ x ∈ v.support, p x) :
793-
(v.subtype_domain p).prod (λi b, h i.1 b) = v.prod h :=
794-
finset.prod_bij (λp _, p.val)
795-
(by { dsimp, simp })
796-
(by simp)
791+
(v.subtype_domain p).prod (λi b, h i b) = v.prod h :=
792+
finset.prod_bij (λp _, p)
793+
(by simp) (by simp)
797794
(assume ⟨a₀, ha₀⟩ ⟨a₁, ha₁⟩, by simp)
798795
(λ i hi, ⟨⟨i, hp i hi⟩, by simpa using hi, rfl⟩)
799796

src/data/equiv/basic.lean

Lines changed: 9 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -588,7 +588,7 @@ def pempty_sum (α : Sort*) : pempty ⊕ α ≃ α :=
588588
@[simp] lemma pempty_sum_apply_inr {α} (a) : pempty_sum α (sum.inr a) = a := rfl
589589

590590
/-- `option α` is equivalent to `α ⊕ punit` -/
591-
def option_equiv_sum_punit (α : Sort*) : option α ≃ α ⊕ punit.{u+1} :=
591+
def option_equiv_sum_punit (α : Type*) : option α ≃ α ⊕ punit.{u+1} :=
592592
⟨λ o, match o with none := inr punit.star | some a := inl a end,
593593
λ s, match s with inr _ := none | inl a := some a end,
594594
λ o, by cases o; refl,
@@ -598,6 +598,14 @@ def option_equiv_sum_punit (α : Sort*) : option α ≃ α ⊕ punit.{u+1} :=
598598
@[simp] lemma option_equiv_sum_punit_some {α} (a) :
599599
option_equiv_sum_punit α (some a) = sum.inl a := rfl
600600

601+
@[simp] lemma option_equiv_sum_punit_symm_inl {α} (a) :
602+
(option_equiv_sum_punit α).symm (sum.inl a) = a :=
603+
rfl
604+
605+
@[simp] lemma option_equiv_sum_punit_symm_inr {α} (a) :
606+
(option_equiv_sum_punit α).symm (sum.inr a) = none :=
607+
rfl
608+
601609
/-- The set of `x : option α` such that `is_some x` is equivalent to `α`. -/
602610
def option_is_some_equiv (α : Type*) : {x : option α // x.is_some} ≃ α :=
603611
{ to_fun := λ o, option.get o.2,

src/data/finset/basic.lean

Lines changed: 6 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -1499,17 +1499,21 @@ protected def subtype {α} (p : α → Prop) [decidable_pred p] (s : finset α)
14991499
λ x y H, subtype.eq $ subtype.mk.inj H⟩
15001500

15011501
@[simp] lemma mem_subtype {p : α → Prop} [decidable_pred p] {s : finset α} :
1502-
∀{a : subtype p}, a ∈ s.subtype p ↔ a.val ∈ s
1502+
∀{a : subtype p}, a ∈ s.subtype p ↔ (a : α) ∈ s
15031503
| ⟨a, ha⟩ := by simp [finset.subtype, ha]
15041504

1505+
lemma subtype_eq_empty {p : α → Prop} [decidable_pred p] {s : finset α} :
1506+
s.subtype p = ∅ ↔ ∀ x, p x → x ∉ s :=
1507+
by simp [ext_iff, subtype.forall, subtype.coe_mk]; refl
1508+
15051509
/-- `s.subtype p` converts back to `s.filter p` with
15061510
`embedding.subtype`. -/
15071511
@[simp] lemma subtype_map (p : α → Prop) [decidable_pred p] :
15081512
(s.subtype p).map (function.embedding.subtype _) = s.filter p :=
15091513
begin
15101514
ext x,
15111515
rw mem_map,
1512-
change (∃ a : {x // p x}, ∃ H, a.val = x) ↔ _,
1516+
change (∃ a : {x // p x}, ∃ H, (a : α) = x) ↔ _,
15131517
split,
15141518
{ rintros ⟨y, hy, hyval⟩,
15151519
rw [mem_subtype, hyval] at hy,

src/data/finsupp/basic.lean

Lines changed: 31 additions & 14 deletions
Original file line numberDiff line numberDiff line change
@@ -209,14 +209,15 @@ by simp only [single_apply]; ac_refl
209209
lemma unique_single [unique α] (x : α →₀ β) : x = single (default α) (x (default α)) :=
210210
by ext i; simp only [unique.eq_default i, single_eq_same]
211211

212+
lemma unique_ext [unique α] {f g : α →₀ β} (h : f (default α) = g (default α)) : f = g :=
213+
ext $ λ a, by rwa [unique.eq_default a]
214+
215+
lemma unique_ext_iff [unique α] {f g : α →₀ β} : f = g ↔ f (default α) = g (default α) :=
216+
⟨λ h, h ▸ rfl, unique_ext⟩
217+
212218
@[simp] lemma unique_single_eq_iff [unique α] {b' : β} :
213219
single a b = single a' b' ↔ b = b' :=
214-
begin
215-
rw [single_eq_single_iff],
216-
split,
217-
{ rintros (⟨_, rfl⟩ | ⟨rfl, rfl⟩); refl },
218-
{ intro h, left, exact ⟨subsingleton.elim _ _, h⟩ }
219-
end
220+
by rw [unique_ext_iff, unique.eq_default a, unique.eq_default a', single_eq_same, single_eq_same]
220221

221222
end single
222223

@@ -309,14 +310,14 @@ begin
309310
{ simp only [h, ne.def, ne_self_iff_false] } }
310311
end
311312

312-
lemma support_emb_domain (f : α₁ ↪ α₂) (v : α₁ →₀ β) :
313+
@[simp] lemma support_emb_domain (f : α₁ ↪ α₂) (v : α₁ →₀ β) :
313314
(emb_domain f v).support = v.support.map f :=
314315
rfl
315316

316-
lemma emb_domain_zero (f : α₁ ↪ α₂) : (emb_domain f 0 : α₂ →₀ β) = 0 :=
317+
@[simp] lemma emb_domain_zero (f : α₁ ↪ α₂) : (emb_domain f 0 : α₂ →₀ β) = 0 :=
317318
rfl
318319

319-
lemma emb_domain_apply (f : α₁ ↪ α₂) (v : α₁ →₀ β) (a : α₁) :
320+
@[simp] lemma emb_domain_apply (f : α₁ ↪ α₂) (v : α₁ →₀ β) (a : α₁) :
320321
emb_domain f v (f a) = v a :=
321322
begin
322323
change dite _ _ _ = _,
@@ -334,10 +335,17 @@ begin
334335
exact set.mem_range_self a
335336
end
336337

337-
lemma emb_domain_inj {f : α₁ ↪ α₂} {l₁ l₂ : α₁ →₀ β} :
338+
lemma emb_domain_injective (f : α₁ ↪ α₂) :
339+
function.injective (emb_domain f : (α₁ →₀ β) → (α₂ →₀ β)) :=
340+
λ l₁ l₂ h, ext $ λ a, by simpa only [emb_domain_apply] using ext_iff.1 h (f a)
341+
342+
@[simp] lemma emb_domain_inj {f : α₁ ↪ α₂} {l₁ l₂ : α₁ →₀ β} :
338343
emb_domain f l₁ = emb_domain f l₂ ↔ l₁ = l₂ :=
339-
⟨λ h, ext $ λ a, by simpa only [emb_domain_apply] using ext_iff.1 h (f a),
340-
λ h, by rw h⟩
344+
(emb_domain_injective f).eq_iff
345+
346+
@[simp] lemma emb_domain_eq_zero {f : α₁ ↪ α₂} {l : α₁ →₀ β} :
347+
emb_domain f l = 0 ↔ l = 0 :=
348+
(emb_domain_injective f).eq_iff' $ emb_domain_zero f
341349

342350
lemma emb_domain_map_range
343351
{β₁ β₂ : Type*} [has_zero β₁] [has_zero β₂]
@@ -1113,7 +1121,7 @@ variables [has_zero β] {v v' : α' →₀ β}
11131121
/-- `subtype_domain p f` is the restriction of the finitely supported function
11141122
`f` to the subtype `p`. -/
11151123
def subtype_domain (p : α → Prop) (f : α →₀ β) : (subtype p →₀ β) :=
1116-
⟨f.support.subtype p, f ∘ subtype.val, λ a, by simp only [mem_subtype, mem_support_iff]⟩
1124+
⟨f.support.subtype p, f ∘ coe, λ a, by simp only [mem_subtype, mem_support_iff]⟩
11171125

11181126
@[simp] lemma support_subtype_domain {f : α →₀ β} :
11191127
(subtype_domain p f).support = f.support.subtype p :=
@@ -1126,10 +1134,19 @@ rfl
11261134
@[simp] lemma subtype_domain_zero : subtype_domain p (0 : α →₀ β) = 0 :=
11271135
rfl
11281136

1137+
lemma subtype_domain_eq_zero_iff' {p : α → Prop} {f : α →₀ β} :
1138+
f.subtype_domain p = 0 ↔ ∀ x, p x → f x = 0 :=
1139+
by simp_rw [← support_eq_empty, support_subtype_domain, subtype_eq_empty, not_mem_support_iff]
1140+
1141+
lemma subtype_domain_eq_zero_iff {p : α → Prop} {f : α →₀ β} (hf : ∀ x ∈ f.support , p x) :
1142+
f.subtype_domain p = 0 ↔ f = 0 :=
1143+
subtype_domain_eq_zero_iff'.trans ⟨λ H, ext $ λ x,
1144+
if hx : p x then H x hx else not_mem_support_iff.1 $ mt (hf x) hx, λ H x _, by simp [H]⟩
1145+
11291146
@[to_additive]
11301147
lemma prod_subtype_domain_index [comm_monoid γ] {v : α →₀ β}
11311148
{h : α → β → γ} (hp : ∀x∈v.support, p x) :
1132-
(v.subtype_domain p).prod (λa b, h a.1 b) = v.prod h :=
1149+
(v.subtype_domain p).prod (λa b, h a b) = v.prod h :=
11331150
prod_bij (λp _, p.val)
11341151
(λ _, mem_subtype.1)
11351152
(λ _ _, rfl)

src/data/option/basic.lean

Lines changed: 10 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -195,4 +195,14 @@ def cases_on' : option α → β → (α → β) → β
195195
| none n s := n
196196
| (some a) n s := s a
197197

198+
@[simp] lemma cases_on'_none (x : β) (f : α → β) : cases_on' none x f = x := rfl
199+
200+
@[simp] lemma cases_on'_some (x : β) (f : α → β) (a : α) : cases_on' (some a) x f = f a := rfl
201+
202+
@[simp] lemma cases_on'_coe (x : β) (f : α → β) (a : α) : cases_on' (a : option α) x f = f a := rfl
203+
204+
@[simp] lemma cases_on'_none_coe (f : option α → β) (o : option α) :
205+
cases_on' o (f none) (f ∘ coe) = f o :=
206+
by cases o; refl
207+
198208
end option

src/data/set/basic.lean

Lines changed: 12 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -693,6 +693,9 @@ by finish [ext_iff, or_comm]
693693
@[simp] theorem pair_eq_singleton (a : α) : ({a, a} : set α) = {a} :=
694694
by finish
695695

696+
theorem pair_comm (a b : α) : ({a, b} : set α) = {b, a} :=
697+
ext $ λ x, or_comm _ _
698+
696699
@[simp] theorem singleton_nonempty (a : α) : ({a} : set α).nonempty :=
697700
⟨a, rfl⟩
698701

@@ -1486,21 +1489,27 @@ by simpa only [exists_prop] using exists_range_iff
14861489
theorem range_iff_surjective : range f = univ ↔ surjective f :=
14871490
eq_univ_iff_forall
14881491

1492+
alias range_iff_surjective ↔ _ function.surjective.range_eq
1493+
14891494
@[simp] theorem range_id : range (@id α) = univ := range_iff_surjective.2 surjective_id
14901495

1496+
theorem is_compl_range_inl_range_inr : is_compl (range $ @sum.inl α β) (range sum.inr) :=
1497+
by { rintro y ⟨⟨x₁, rfl⟩, ⟨x₂, _⟩⟩, cc },
1498+
by { rintro (x|y) -; [left, right]; exact mem_range_self _ }⟩
1499+
14911500
theorem range_inl_union_range_inr : range (@sum.inl α β) ∪ range sum.inr = univ :=
1492-
by { ext x, cases x; simp }
1501+
is_compl_range_inl_range_inr.sup_eq_top
14931502

14941503
@[simp] theorem range_quot_mk (r : α → α → Prop) : range (quot.mk r) = univ :=
1495-
range_iff_surjective.2 quot.exists_rep
1504+
(surjective_quot_mk r).range_eq
14961505

14971506
@[simp] theorem image_univ {ι : Type*} {f : ι → β} : f '' univ = range f :=
14981507
by { ext, simp [image, range] }
14991508

15001509
theorem image_subset_range {ι : Type*} (f : ι → β) (s : set ι) : f '' s ⊆ range f :=
15011510
by rw ← image_univ; exact image_subset _ (subset_univ _)
15021511

1503-
theorem range_comp {g : α → β} : range (g ∘ f) = g '' range f :=
1512+
theorem range_comp (g : α → β) (f : ι → α) : range (g ∘ f) = g '' range f :=
15041513
subset.antisymm
15051514
(forall_range_iff.mpr $ assume i, mem_image_of_mem g (mem_range_self _))
15061515
(ball_image_iff.mpr $ forall_range_iff.mpr mem_range_self)
@@ -1699,9 +1708,6 @@ by { intro s, use f ⁻¹' s, rw image_preimage_eq hf }
16991708
lemma injective.image_injective (hf : injective f) : injective (image f) :=
17001709
by { intros s t h, rw [←preimage_image_eq s hf, ←preimage_image_eq t hf, h] }
17011710

1702-
lemma surjective.range_eq {f : ι → α} (hf : surjective f) : range f = univ :=
1703-
range_iff_surjective.2 hf
1704-
17051711
lemma surjective.preimage_subset_preimage_iff {s t : set β} (hf : surjective f) :
17061712
f ⁻¹' s ⊆ f ⁻¹' t ↔ s ⊆ t :=
17071713
by { apply preimage_subset_preimage_iff, rw [hf.range_eq], apply subset_univ }

src/data/set/function.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -49,7 +49,7 @@ lemma restrict_eq (f : α → β) (s : set α) : s.restrict f = f ∘ coe := rfl
4949
@[simp] lemma restrict_apply (f : α → β) (s : set α) (x : s) : restrict f s x = f x := rfl
5050

5151
@[simp] lemma range_restrict (f : α → β) (s : set α) : set.range (restrict f s) = f '' s :=
52-
range_comp.trans $ congr_arg (('') f) subtype.range_coe
52+
(range_comp _ _).trans $ congr_arg (('') f) subtype.range_coe
5353

5454
/-- Restrict codomain of a function `f` to a set `s`. Same as `subtype.coind` but this version
5555
has codomain `↥s` instead of `subtype s`. -/

src/data/set/lattice.lean

Lines changed: 3 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -104,6 +104,9 @@ theorem subset_Inter {t : set β} {s : ι → set β} (h : ∀ i, t ⊆ s i) : t
104104
-- TODO: should be simpler when sets' order is based on lattices
105105
@le_infi (set β) _ set.lattice_set _ _ h
106106

107+
theorem subset_Inter_iff {t : set β} {s : ι → set β} : t ⊆ (⋂ i, s i) ↔ ∀ i, t ⊆ s i :=
108+
@le_infi_iff (set β) _ set.lattice_set _ _
109+
107110
theorem subset_Union : ∀ (s : ι → set β) (i : ι), s i ⊆ (⋃ i, s i) := le_supr
108111

109112
-- This rather trivial consequence is convenient with `apply`,

src/data/sum.lean

Lines changed: 6 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -48,6 +48,12 @@ end⟩
4848

4949
namespace sum
5050

51+
lemma injective_inl : function.injective (sum.inl : α → α ⊕ β) :=
52+
λ x y, sum.inl.inj
53+
54+
lemma injective_inr : function.injective (sum.inr : β → α ⊕ β) :=
55+
λ x y, sum.inr.inj
56+
5157
/-- Map `α ⊕ β` to `α' ⊕ β'` sending `α` to `α'` and `β` to `β'`. -/
5258
protected def map (f : α → α') (g : β → β') : α ⊕ β → α' ⊕ β'
5359
| (sum.inl x) := sum.inl (f x)

src/logic/nontrivial.lean

Lines changed: 3 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -70,6 +70,9 @@ lemma subsingleton_iff : subsingleton α ↔ ∀ (x y : α), x = y :=
7070
lemma not_nontrivial_iff_subsingleton : ¬(nontrivial α) ↔ subsingleton α :=
7171
by { rw [nontrivial_iff, subsingleton_iff], push_neg, refl }
7272

73+
lemma not_subsingleton (α) [h : nontrivial α] : ¬subsingleton α :=
74+
let ⟨⟨x, y, hxy⟩⟩ := h in λ ⟨h'⟩, hxy $ h' x y
75+
7376
/-- A type is either a subsingleton or nontrivial. -/
7477
lemma subsingleton_or_nontrivial (α : Type*) : subsingleton α ∨ nontrivial α :=
7578
by { rw [← not_nontrivial_iff_subsingleton, or_comm], exact classical.em _ }

0 commit comments

Comments
 (0)