refactor(library/data/finset/basic,library/*): get rid of finset singleton

This commit is contained in:
Jeremy Avigad 2015-12-31 11:50:55 -08:00 committed by Leonardo de Moura
parent 8a00a431e8
commit 12a69bad04
6 changed files with 50 additions and 64 deletions
library
data/finset
theories

View file

@ -83,22 +83,6 @@ theorem mem_of_mem_list {a : A} {l : nodup_list A} : a ∈ elt_of l → a ∈
theorem mem_list_of_mem {a : A} {l : nodup_list A} : a ∈ ⟦l⟧ → a ∈ elt_of l :=
λ ainl, ainl
/- singleton -/
definition singleton (a : A) : finset A :=
to_finset_of_nodup [a] !nodup_singleton
theorem mem_singleton [simp] (a : A) : a ∈ singleton a :=
mem_of_mem_list !mem_cons
theorem eq_of_mem_singleton {x a : A} : x ∈ singleton a → x = a :=
list.mem_singleton
theorem mem_singleton_eq (x a : A) : (x ∈ singleton a) = (x = a) :=
propext (iff.intro eq_of_mem_singleton (assume H, eq.subst H !mem_singleton))
lemma eq_of_singleton_eq {a b : A} : singleton a = singleton b → a = b :=
assume Pseq, eq_of_mem_singleton (Pseq ▸ mem_singleton a)
definition decidable_mem [instance] [h : decidable_eq A] : ∀ (a : A) (s : finset A), decidable (a ∈ s) :=
λ a s, quot.rec_on_subsingleton s
(λ l, match list.decidable_mem a (elt_of l) with
@ -152,9 +136,6 @@ quot.lift_on s
theorem card_empty : card (@empty A) = 0 :=
rfl
theorem card_singleton (a : A) : card (singleton a) = 1 :=
rfl
lemma ne_empty_of_card_eq_succ {s : finset A} {n : nat} : card s = succ n → s ≠ ∅ :=
by intros; substvars; contradiction
@ -185,15 +166,27 @@ quot.induction_on s (λ l : nodup_list A, λ H, list.eq_or_mem_of_mem_insert H)
theorem mem_of_mem_insert_of_ne {x a : A} {s : finset A} (xin : x ∈ insert a s) : x ≠ a → x ∈ s :=
or_resolve_right (eq_or_mem_of_mem_insert xin)
theorem mem_insert_iff (x a : A) (s : finset A) : x ∈ insert a s ↔ (x = a x ∈ s) :=
iff.intro !eq_or_mem_of_mem_insert
(or.rec (λH', (eq.substr H' !mem_insert)) !mem_insert_of_mem)
theorem mem_insert_eq (x a : A) (s : finset A) : x ∈ insert a s = (x = a x ∈ s) :=
propext (iff.intro !eq_or_mem_of_mem_insert
(or.rec (λH', (eq.substr H' !mem_insert)) !mem_insert_of_mem))
propext !mem_insert_iff
theorem insert_empty_eq (a : A) : '{a} = singleton a := rfl
theorem mem_singleton_eq' (x a : A) : x ∈ '{a} = (x = a) :=
theorem mem_singleton_iff (x a : A) : x ∈ '{a} ↔ (x = a) :=
by rewrite [mem_insert_eq, mem_empty_eq, or_false]
theorem mem_singleton (a : A) : a ∈ '{a} := mem_insert a ∅
theorem mem_singleton_of_eq {x a : A} (H : x = a) : x ∈ '{a} :=
by rewrite H; apply mem_insert
theorem eq_of_mem_singleton {x a : A} (H : x ∈ '{a}) : x = a := iff.mp !mem_singleton_iff H
theorem eq_of_singleton_eq {a b : A} (H : '{a} = '{b}) : a = b :=
have a ∈ '{b}, by rewrite -H; apply mem_singleton,
eq_of_mem_singleton this
theorem insert_eq_of_mem {a : A} {s : finset A} (H : a ∈ s) : insert a s = s :=
ext (λ x, eq.substr (mem_insert_eq x a s)
(or_iff_right_of_imp (λH1, eq.substr H1 H)))
@ -396,16 +389,11 @@ theorem empty_union (s : finset A) : ∅ s = s :=
calc ∅ s = s ∅ : union.comm
... = s : union_empty
theorem insert_eq (a : A) (s : finset A) : insert a s = singleton a s :=
ext (take x,
calc
x ∈ insert a s ↔ x ∈ insert a s : iff.refl
... = (x = a x ∈ s) : mem_insert_eq
... = (x ∈ singleton a x ∈ s) : mem_singleton_eq
... = (x ∈ '{a} s) : mem_union_eq)
theorem insert_eq (a : A) (s : finset A) : insert a s = '{a} s :=
ext (take x, by rewrite [mem_insert_iff, mem_union_iff, mem_singleton_iff])
theorem insert_union (a : A) (s t : finset A) : insert a (s t) = insert a s t :=
by rewrite [*insert_eq, union.assoc]
by rewrite [insert_eq, insert_eq a s, union.assoc]
end union
/- inter -/
@ -466,20 +454,20 @@ calc ∅ ∩ s = s ∩ ∅ : inter.comm
... = ∅ : inter_empty
theorem singleton_inter_of_mem {a : A} {s : finset A} (H : a ∈ s) :
singleton a ∩ s = singleton a :=
'{a} ∩ s = '{a} :=
ext (take x,
begin
rewrite [mem_inter_eq, !mem_singleton_eq],
rewrite [mem_inter_eq, !mem_singleton_iff],
exact iff.intro
(suppose x = a ∧ x ∈ s, and.left this)
(suppose x = a, and.intro this (eq.subst (eq.symm this) H))
end)
theorem singleton_inter_of_not_mem {a : A} {s : finset A} (H : a ∉ s) :
singleton a ∩ s = ∅ :=
'{a} ∩ s = ∅ :=
ext (take x,
begin
rewrite [mem_inter_eq, !mem_singleton_eq, mem_empty_eq],
rewrite [mem_inter_eq, !mem_singleton_iff, mem_empty_eq],
exact iff.intro
(suppose x = a ∧ x ∈ s, H (eq.subst (and.left this) (and.right this)))
(false.elim)
@ -676,7 +664,7 @@ theorem upto_zero : upto 0 = ∅ := rfl
theorem upto_succ (n : ) : upto (succ n) = upto n '{n} :=
begin
apply ext, intro x,
rewrite [mem_union_iff, *mem_upto_iff, mem_singleton_eq', lt_succ_iff_le, nat.le_iff_lt_or_eq],
rewrite [mem_union_iff, *mem_upto_iff, mem_singleton_iff, lt_succ_iff_le, nat.le_iff_lt_or_eq],
end
/- useful rules for calculations with quantifiers -/

View file

@ -441,7 +441,7 @@ begin
induction s with a s nains ih,
intro x,
rewrite powerset_empty,
show x ∈ '{∅} ↔ x ⊆ ∅, by rewrite [mem_singleton_eq', subset_empty_iff],
show x ∈ '{∅} ↔ x ⊆ ∅, by rewrite [mem_singleton_iff, subset_empty_iff],
intro x,
rewrite [powerset_insert nains, mem_union_iff, ih, mem_image_iff],
exact

View file

@ -130,15 +130,15 @@ private theorem aux₀ (s : finset A) : {t ∈ powerset s | card t = 0} = '{∅}
ext (take t, iff.intro
(assume H,
assert t = ∅, from eq_empty_of_card_eq_zero (of_mem_sep H),
show t ∈ '{ ∅ }, by rewrite [this, mem_singleton_eq'])
show t ∈ '{ ∅ }, by rewrite [this, mem_singleton_iff])
(assume H,
assert t = ∅, by rewrite mem_singleton_eq' at H; assumption,
assert t = ∅, by rewrite mem_singleton_iff at H; assumption,
by substvars; exact mem_sep_of_mem !empty_mem_powerset rfl))
private theorem aux₁ (k : ) : {t ∈ powerset (∅ : finset A) | card t = succ k} = ∅ :=
eq_empty_of_forall_not_mem (take t, assume H,
assert t ∈ powerset ∅, from mem_of_mem_sep H,
assert t = ∅, by rewrite [powerset_empty at this, mem_singleton_eq' at this]; assumption,
assert t = ∅, by rewrite [powerset_empty at this, mem_singleton_iff at this]; assumption,
have card (∅ : finset A) = succ k, by rewrite -this; apply of_mem_sep H,
nat.no_confusion this)

View file

@ -26,7 +26,7 @@ definition orbit (hom : G → perm S) (H : finset G) (a : S) : finset S :=
image (move_by a) (image hom H)
definition fixed_points [reducible] (hom : G → perm S) (H : finset G) : finset S :=
{a ∈ univ | orbit hom H a = singleton a}
{a ∈ univ | orbit hom H a = '{a}}
variable [decidable_eq G] -- required by {x ∈ H |p x} filtering
@ -40,7 +40,7 @@ end def
section orbit_stabilizer
variables {G S : Type} [group G] [fintype S] [decidable_eq S]
variables {G S : Type} [group G] [decidable_eq G] [fintype S] [decidable_eq S]
section
@ -71,9 +71,9 @@ lemma mem_fixed_points_of_exists_of_is_fixed_point :
assume Pex Pfp, mem_sep_of_mem !mem_univ
(ext take x, iff.intro
(assume Porb, obtain h Phin Pha, from exists_of_orbit Porb,
by rewrite [mem_singleton_eq, -Pha, Pfp h Phin])
by rewrite [mem_singleton_iff, -Pha, Pfp h Phin])
(obtain h Phin, from Pex,
by rewrite mem_singleton_eq;
by rewrite mem_singleton_iff;
intro Peq; rewrite Peq;
apply orbit_of_exists;
existsi h; apply and.intro Phin (Pfp h Phin)))
@ -86,22 +86,20 @@ lemma is_fixed_point_iff_mem_fixed_points [finsubgH : is_finsubg H] :
a ∈ fixed_points hom H ↔ is_fixed_point hom H a :=
is_fixed_point_iff_mem_fixed_points_of_exists (exists.intro 1 !finsubg_has_one)
lemma is_fixed_point_of_one : is_fixed_point hom (singleton 1) a :=
lemma is_fixed_point_of_one : is_fixed_point hom ('{1}) a :=
take h, assume Ph, by rewrite [eq_of_mem_singleton Ph, hom_map_one]
lemma fixed_points_of_one : fixed_points hom (singleton 1) = univ :=
lemma fixed_points_of_one : fixed_points hom ('{1}) = univ :=
ext take s, iff.intro (assume Pl, mem_univ s)
(assume Pr, mem_fixed_points_of_exists_of_is_fixed_point
(exists.intro 1 !mem_singleton) is_fixed_point_of_one)
open fintype
lemma card_fixed_points_of_one : card (fixed_points hom (singleton 1)) = card S :=
lemma card_fixed_points_of_one : card (fixed_points hom ('{1})) = card S :=
by rewrite [fixed_points_of_one]
end
variable [decidable_eq G]
-- these are already specified by stab hom H a
variables {hom : G → perm S} {H : finset G} {a : S}
@ -143,7 +141,8 @@ lemma moverset_inj_on_orbit : set.inj_on (moverset hom H a) (ts (orbit hom H a))
apply of_mem_sep Ph1b1
end
variable [is_finsubg H]
variable [finsubgH : is_finsubg H]
include finsubgH
lemma subg_stab_of_move {h g : G} :
h ∈ H → g ∈ moverset hom H a (hom h a) → h⁻¹*g ∈ stab hom H a :=
@ -250,7 +249,7 @@ end orbit_stabilizer
section orbit_partition
variables {G S : Type} [group G] [fintype S] [decidable_eq S]
variables {G S : Type} [group G] [decidable_eq G] [fintype S] [decidable_eq S]
variables {hom : G → perm S} [Hom : is_hom_class hom] {H : finset G} [subgH : is_finsubg H]
include Hom subgH
@ -325,7 +324,7 @@ ext take s, iff.intro
obtain a Pa, from exists_of_mem_orbits Psin,
mem_image
(mem_sep_of_mem !mem_univ (eq.symm
(eq_of_card_eq_of_subset (by rewrite [card_singleton, Pa, Ps])
(eq_of_card_eq_of_subset (by rewrite [Pa, Ps])
(subset_of_forall
take x, assume Pxin, eq_of_mem_singleton Pxin ▸ in_orbit_refl))))
Pa)

View file

@ -19,6 +19,7 @@ open ops
section subg
-- we should be able to prove properties using finsets directly
variables {G : Type} [group G]
variable [decidable_eq G]
definition finset_mul_closed_on [reducible] (H : finset G) : Prop :=
∀ x y : G, x ∈ H → y ∈ H → x * y ∈ H
@ -33,19 +34,17 @@ structure is_finsubg [class] (H : finset G) : Type :=
definition univ_is_finsubg [instance] [finG : fintype G] : is_finsubg (@finset.univ G _) :=
is_finsubg.mk !mem_univ (λ x y Px Py, !mem_univ) (λ a Pa, !mem_univ)
definition one_is_finsubg [instance] : is_finsubg (singleton (1:G)) :=
definition one_is_finsubg [instance] : is_finsubg ('{(1:G)}) :=
is_finsubg.mk !mem_singleton
(λ x y Px Py, by rewrite [eq_of_mem_singleton Px, eq_of_mem_singleton Py, one_mul]; apply mem_singleton)
(λ x Px, by rewrite [eq_of_mem_singleton Px, one_inv]; apply mem_singleton)
lemma finsubg_has_one (H : finset G) [h : is_finsubg H] : 1 ∈ H :=
@is_finsubg.has_one G _ H h
@is_finsubg.has_one G _ _ H h
lemma finsubg_mul_closed (H : finset G) [h : is_finsubg H] {x y : G} : x ∈ H → y ∈ H → x * y ∈ H :=
@is_finsubg.mul_closed G _ H h x y
@is_finsubg.mul_closed G _ _ H h x y
lemma finsubg_has_inv (H : finset G) [h : is_finsubg H] {a : G} : a ∈ H → a⁻¹ ∈ H :=
@is_finsubg.has_inv G _ H h a
variable [decidable_eq G]
@is_finsubg.has_inv G _ _ H h a
definition finsubg_to_subg [instance] {H : finset G} [h : is_finsubg H]
: is_subgroup (ts H) :=
@ -58,10 +57,10 @@ definition finsubg_to_subg [instance] {H : finset G} [h : is_finsubg H]
open nat
lemma finsubg_eq_singleton_one_of_card_one {H : finset G} [h : is_finsubg H] :
card H = 1 → H = singleton 1 :=
card H = 1 → H = '{1} :=
assume Pcard, eq.symm (eq_of_card_eq_of_subset (by rewrite [Pcard])
(subset_of_forall take g,
by rewrite [mem_singleton_eq]; intro Pg; rewrite Pg; exact finsubg_has_one H))
by rewrite [mem_singleton_iff]; intro Pg; rewrite Pg; exact finsubg_has_one H))
end subg

View file

@ -32,7 +32,7 @@ lemma card_mod_eq_of_action_by_psubg {p : nat} :
| 0 := by rewrite [↑psubg, pow_zero]; intro Psubg;
rewrite [finsubg_eq_singleton_one_of_card_one (and.right Psubg), fixed_points_of_one]
| (succ m) := take Ppsubg, begin
rewrite [@orbit_class_equation' G S ambientG finS deceqS hom Hom H subgH],
rewrite [@orbit_class_equation' G S ambientG _ finS deceqS hom Hom H subgH],
apply add_mod_eq_of_dvd, apply dvd_Sum_of_dvd,
intro s Psin,
rewrite mem_sep_iff at Psin,
@ -370,9 +370,9 @@ include ambientG deceqG finG
theorem first_sylow_theorem {p : nat} (Pp : prime p) :
∀ n, p^n card G → ∃ (H : finset G) (finsubgH : is_finsubg H), card H = p^n
| 0 := assume Pdvd, exists.intro (singleton 1)
| 0 := assume Pdvd, exists.intro ('{1})
(exists.intro one_is_finsubg
(by rewrite [card_singleton, pow_zero]))
(by rewrite [pow_zero]))
| (succ n) := assume Pdvd,
obtain H PfinsubgH PcardH, from first_sylow_theorem n (pow_dvd_of_pow_succ_dvd Pdvd),
assert Ppsubg : psubg H p n, from and.intro Pp PcardH,