refactor(library/data): cleanup proofs using new features

This commit is contained in:
Leonardo de Moura 2015-07-21 08:26:44 -07:00
parent 9a85a95893
commit 3e3d37905c
3 changed files with 96 additions and 96 deletions

View file

@ -183,9 +183,9 @@ theorem mem_of_mem_insert_of_ne {x a : A} {s : finset A} : x ∈ insert a s →
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)
(assume H, or.elim H
(assume H' : x = a, eq.subst (eq.symm H') !mem_insert)
(assume H' : x ∈ s, !mem_insert_of_mem H')))
(suppose x = a x ∈ s, or.elim this
(suppose x = a, eq.subst (eq.symm this) !mem_insert)
(suppose x ∈ s, !mem_insert_of_mem this)))
theorem insert_empty_eq (a : A) : '{a} = singleton a := rfl
@ -197,9 +197,9 @@ ext
show x = a x ∈ s ↔ x ∈ s, from
iff.intro
(assume H1, or.elim H1
(assume H2 : x = a, eq.subst (eq.symm H2) H)
(assume H2, H2))
(assume H1, or.inr H1)
(suppose x = a, eq.subst (eq.symm this) H)
(suppose x ∈ s, this))
(suppose x ∈ s, or.inr this)
end
theorem card_insert_of_mem {a : A} {s : finset A} : a ∈ s → card (insert a s) = card s :=
@ -213,8 +213,8 @@ quot.induction_on s
theorem card_insert_le (a : A) (s : finset A) :
card (insert a s) ≤ card s + 1 :=
decidable.by_cases
(assume H : a ∈ s, by rewrite [card_insert_of_mem H]; apply le_succ)
(assume H : a ∉ s, by rewrite [card_insert_of_not_mem H])
(suppose a ∈ s, by rewrite [card_insert_of_mem this]; apply le_succ)
(suppose a ∉ s, by rewrite [card_insert_of_not_mem this])
lemma ne_empty_of_card_eq_succ {s : finset A} {n : nat} : card s = succ n → s ≠ ∅ :=
by intros; substvars; contradiction
@ -232,16 +232,16 @@ quot.induction_on s
(assume nodup_l, H1)
(take a l',
assume IH nodup_al',
have anl' : a ∉ l', from not_mem_of_nodup_cons nodup_al',
assert H3 : list.insert a l' = a :: l', from insert_eq_of_not_mem anl',
assert nodup_l' : nodup l', from nodup_of_nodup_cons nodup_al',
assert P_l' : P (quot.mk (subtype.tag l' nodup_l')), from IH nodup_l',
assert H4 : P (insert a (quot.mk (subtype.tag l' nodup_l'))), from H2 anl' P_l',
have a ∉ l', from not_mem_of_nodup_cons nodup_al',
assert e : list.insert a l' = a :: l', from insert_eq_of_not_mem this,
assert nodup l', from nodup_of_nodup_cons nodup_al',
assert P (quot.mk (subtype.tag l' this)), from IH this,
assert P (insert a (quot.mk (subtype.tag l' _))), from H2 `a ∉ l'` this,
begin
revert nodup_al',
rewrite [-H3],
rewrite [-e],
intros,
apply H4
apply this
end)))
protected theorem induction_on {P : finset A → Prop} (s : finset A)
@ -306,20 +306,20 @@ theorem erase_insert (a : A) (s : finset A) : a ∉ s → erase a (insert a s) =
(λ bin, by subst b; contradiction))
(λ bnea : b ≠ a, iff.intro
(λ bin,
assert bin' : b ∈ insert a s, from mem_of_mem_erase bin,
mem_of_mem_insert_of_ne bin' bnea)
assert b ∈ insert a s, from mem_of_mem_erase bin,
mem_of_mem_insert_of_ne this bnea)
(λ bin,
have bin' : b ∈ insert a s, from mem_insert_of_mem _ bin,
mem_erase_of_ne_of_mem bnea bin')))
have b ∈ insert a s, from mem_insert_of_mem _ bin,
mem_erase_of_ne_of_mem bnea this)))
theorem insert_erase {a : A} {s : finset A} : a ∈ s → insert a (erase a s) = s :=
λ ains, finset.ext (λ b, by_cases
(λ beqa : b = a, iff.intro
(suppose b = a, iff.intro
(λ bin, by subst b; assumption)
(λ bin, by subst b; apply mem_insert))
(λ bnea : b ≠ a, iff.intro
(λ bin, mem_of_mem_erase (mem_of_mem_insert_of_ne bin bnea))
(λ bin, mem_insert_of_mem _ (mem_erase_of_ne_of_mem bnea bin))))
(suppose b ≠ a, iff.intro
(λ bin, mem_of_mem_erase (mem_of_mem_insert_of_ne bin `b ≠ a`))
(λ bin, mem_insert_of_mem _ (mem_erase_of_ne_of_mem `b ≠ a` bin))))
end erase
/- union -/
@ -374,8 +374,8 @@ ext (λ a, iff.intro
theorem union_empty (s : finset A) : s ∅ = s :=
ext (λ a, iff.intro
(λ ain : a ∈ s ∅, or.elim (mem_or_mem_of_mem_union ain) (λ i, i) (λ i, absurd i !not_mem_empty))
(λ i : a ∈ s, mem_union_left _ i))
(suppose a ∈ s ∅, or.elim (mem_or_mem_of_mem_union this) (λ i, i) (λ i, absurd i !not_mem_empty))
(suppose a ∈ s, mem_union_left _ this))
theorem empty_union (s : finset A) : ∅ s = s :=
calc ∅ s = s ∅ : union.comm
@ -437,8 +437,8 @@ ext (λ a, iff.intro
theorem inter_empty (s : finset A) : s ∩ ∅ = ∅ :=
ext (λ a, iff.intro
(λ h : a ∈ s ∩ ∅, absurd (mem_of_mem_inter_right h) !not_mem_empty)
(λ h : a ∈ ∅, absurd h !not_mem_empty))
(suppose a ∈ s ∩ ∅, absurd (mem_of_mem_inter_right this) !not_mem_empty)
(suppose a ∈ ∅, absurd this !not_mem_empty))
theorem empty_inter (s : finset A) : ∅ ∩ s = ∅ :=
calc ∅ ∩ s = s ∩ ∅ : inter.comm
@ -450,8 +450,8 @@ ext (take x,
begin
rewrite [mem_inter_eq, !mem_singleton_eq],
exact iff.intro
(assume H1 : x = a ∧ x ∈ s, and.left H1)
(assume H1 : x = a, and.intro H1 (eq.subst (eq.symm H1) H))
(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) :
@ -460,7 +460,7 @@ ext (take x,
begin
rewrite [mem_inter_eq, !mem_singleton_eq, mem_empty_eq],
exact iff.intro
(assume H1 : x = a ∧ x ∈ s, H (eq.subst (and.left H1) (and.right H1)))
(suppose x = a ∧ x ∈ s, H (eq.subst (and.left this) (and.right this)))
(false.elim)
end)
end inter
@ -489,13 +489,13 @@ definition disjoint (s₁ s₂ : finset A) : Prop :=
quot.lift_on₂ s₁ s₂ (λ l₁ l₂, disjoint (elt_of l₁) (elt_of l₂))
(λ v₁ v₂ w₁ w₂ p₁ p₂, propext (iff.intro
(λ d₁ a (ainw₁ : a ∈ elt_of w₁),
have ainv₁ : a ∈ elt_of v₁, from mem_perm (perm.symm p₁) ainw₁,
have nainv₂ : a ∉ elt_of v₂, from disjoint_left d₁ ainv₁,
not_mem_perm p₂ nainv₂)
have a ∈ elt_of v₁, from mem_perm (perm.symm p₁) ainw₁,
have a ∉ elt_of v₂, from disjoint_left d₁ this,
not_mem_perm p₂ this)
(λ d₂ a (ainv₁ : a ∈ elt_of v₁),
have ainw₁ : a ∈ elt_of w₁, from mem_perm p₁ ainv₁,
have nainw₂ : a ∉ elt_of w₂, from disjoint_left d₂ ainw₁,
not_mem_perm (perm.symm p₂) nainw₂)))
have a ∈ elt_of w₁, from mem_perm p₁ ainv₁,
have a ∉ elt_of w₂, from disjoint_left d₂ this,
not_mem_perm (perm.symm p₂) this)))
theorem disjoint.elim {s₁ s₂ : finset A} {x : A} :
disjoint s₁ s₂ → x ∈ s₁ → x ∈ s₂ → false :=
@ -510,8 +510,8 @@ ext (take x, iff_false_intro (assume H1,
theorem disjoint_of_inter_eq_empty [h : decidable_eq A] {s₁ s₂ : finset A} (H : s₁ ∩ s₂ = ∅) : disjoint s₁ s₂ :=
disjoint.intro (take x H1 H2,
have H3 : x ∈ s₁ ∩ s₂, from mem_inter H1 H2,
!not_mem_empty (eq.subst H H3))
have x ∈ s₁ ∩ s₂, from mem_inter H1 H2,
!not_mem_empty (eq.subst H this))
theorem disjoint.comm {s₁ s₂ : finset A} : disjoint s₁ s₂ → disjoint s₂ s₁ :=
quot.induction_on₂ s₁ s₂ (λ l₁ l₂ d, list.disjoint.comm d)
@ -549,7 +549,7 @@ theorem subset_of_forall {s₁ s₂ : finset A} : (∀x, x ∈ s₁ → x ∈ s
quot.induction_on₂ s₁ s₂ (λ l₁ l₂ H, H)
theorem subset_insert [h : decidable_eq A] (s : finset A) (a : A) : s ⊆ insert a s :=
subset_of_forall (take x, assume H : x ∈ s, mem_insert_of_mem _ H)
subset_of_forall (take x, suppose x ∈ s, mem_insert_of_mem _ this)
theorem eq_of_subset_of_subset {s₁ s₂ : finset A} (H₁ : s₁ ⊆ s₂) (H₂ : s₂ ⊆ s₁) : s₁ = s₂ :=
ext (take x, iff.intro (assume H, mem_of_subset_of_mem H₁ H) (assume H, mem_of_subset_of_mem H₂ H))
@ -605,13 +605,13 @@ iff.intro
(assume H,
obtain x [H1 H2], from H,
or.elim (eq_or_mem_of_mem_insert H1)
(assume H3 : x = a, or.inl (eq.subst H3 H2))
(assume H3 : x ∈ s, or.inr (exists.intro x (and.intro H3 H2))))
(suppose x = a, or.inl (eq.subst this H2))
(suppose x ∈ s, or.inr (exists.intro x (and.intro this H2))))
(assume H,
or.elim H
(assume H1 : P a, exists.intro a (and.intro !mem_insert H1))
(assume H1 : ∃ x, x ∈ s ∧ P x,
obtain x [H2 H3], from H1,
(suppose P a, exists.intro a (and.intro !mem_insert this))
(suppose ∃ x, x ∈ s ∧ P x,
obtain x [H2 H3], from this,
exists.intro x (and.intro (!mem_insert_of_mem H2) H3)))
theorem exists_mem_insert_eq {A : Type} [d : decidable_eq A] (a : A) (s : finset A) (P : A → Prop) :
@ -633,8 +633,8 @@ iff.intro
(assume H, and.intro (H _ !mem_insert) (take x, assume H', H _ (!mem_insert_of_mem H')))
(assume H, take x, assume H' : x ∈ insert a s,
or.elim (eq_or_mem_of_mem_insert H')
(assume H1 : x = a, eq.subst (eq.symm H1) (and.left H))
(assume H1 : x ∈ s, and.right H _ H1))
(suppose x = a, eq.subst (eq.symm this) (and.left H))
(suppose x ∈ s, and.right H _ this))
theorem forall_mem_insert_eq {A : Type} [d : decidable_eq A] (a : A) (s : finset A) (P : A → Prop) :
(∀ x, x ∈ insert a s → P x) = (P a ∧ (∀ x, x ∈ s → P x)) :=

View file

@ -60,29 +60,29 @@ ext (take y, iff.intro
obtain x (H1l : x ∈ insert a s) (H1r :f x = y), from exists_of_mem_image H,
have H2 : x = a x ∈ s, from eq_or_mem_of_mem_insert H1l,
or.elim H2
(assume H3 : x = a,
have H4 : f a = y, from eq.subst H3 H1r,
show y ∈ insert (f a) (image f s), from eq.subst H4 !mem_insert)
(assume H3 : x ∈ s,
have H5 : f x ∈ image f s, from mem_image_of_mem f H3,
show y ∈ insert (f a) (image f s), from eq.subst H1r (mem_insert_of_mem _ H5)))
(assume H : y ∈ insert (f a) (image f s),
have H1 : y = f a y ∈ image f s, from eq_or_mem_of_mem_insert H,
or.elim H1
(assume H2 : y = f a,
have H3 : f a ∈ image f (insert a s), from mem_image_of_mem f !mem_insert,
show y ∈ image f (insert a s), from eq.subst (eq.symm H2) H3)
(assume H2 : y ∈ image f s,
show y ∈ image f (insert a s), from mem_image_of_mem_image_of_subset H2 !subset_insert)))
(suppose x = a,
have f a = y, from eq.subst this H1r,
show y ∈ insert (f a) (image f s), from eq.subst this !mem_insert)
(suppose x ∈ s,
have f x ∈ image f s, from mem_image_of_mem f this,
show y ∈ insert (f a) (image f s), from eq.subst H1r (mem_insert_of_mem _ this)))
(suppose y ∈ insert (f a) (image f s),
have y = f a y ∈ image f s, from eq_or_mem_of_mem_insert this,
or.elim this
(suppose y = f a,
have f a ∈ image f (insert a s), from mem_image_of_mem f !mem_insert,
show y ∈ image f (insert a s), from eq.subst (eq.symm `y = f a`) this)
(suppose y ∈ image f s,
show y ∈ image f (insert a s), from mem_image_of_mem_image_of_subset this !subset_insert)))
lemma image_compose {C : Type} [deceqC : decidable_eq C] {f : B → C} {g : A → B} {s : finset A} :
image (f∘g) s = image f (image g s) :=
ext (take z, iff.intro
(assume Hz : z ∈ image (f∘g) s,
obtain x (Hx : x ∈ s) (Hgfx : f (g x) = z), from exists_of_mem_image Hz,
(suppose z ∈ image (f∘g) s,
obtain x (Hx : x ∈ s) (Hgfx : f (g x) = z), from exists_of_mem_image this,
by rewrite -Hgfx; apply mem_image_of_mem _ (mem_image_of_mem _ Hx))
(assume Hz : z ∈ image f (image g s),
obtain y (Hy : y ∈ image g s) (Hfy : f y = z), from exists_of_mem_image Hz,
(suppose z ∈ image f (image g s),
obtain y (Hy : y ∈ image g s) (Hfy : f y = z), from exists_of_mem_image this,
obtain x (Hx : x ∈ s) (Hgx : g x = y), from exists_of_mem_image Hy,
mem_image_of_mem_of_eq Hx (by esimp; rewrite [Hgx, Hfy])))
@ -155,14 +155,14 @@ propext !mem_diff_iff
theorem union_diff_cancel {s t : finset A} (H : s ⊆ t) : s (t \ s) = t :=
ext (take x, iff.intro
(assume H1 : x ∈ s (t \ s),
or.elim (mem_or_mem_of_mem_union H1)
(assume H2 : x ∈ s, mem_of_subset_of_mem H H2)
(assume H2 : x ∈ t \ s, mem_of_mem_diff H2))
(assume H1 : x ∈ t,
(suppose x ∈ s (t \ s),
or.elim (mem_or_mem_of_mem_union this)
(suppose x ∈ s, mem_of_subset_of_mem H this)
(suppose x ∈ t \ s, mem_of_mem_diff this))
(suppose x ∈ t,
decidable.by_cases
(assume H2 : x ∈ s, mem_union_left _ H2)
(assume H2 : x ∉ s, mem_union_right _ (mem_diff H1 H2))))
(suppose x ∈ s, mem_union_left _ this)
(suppose x ∉ s, mem_union_right _ (mem_diff `x ∈ t` this))))
theorem diff_union_cancel {s t : finset A} (H : s ⊆ t) : (t \ s) s = t :=
eq.subst !union.comm (!union_diff_cancel H)
@ -224,8 +224,8 @@ quot.induction_on₂ s₁ s₂ (λ l₁ l₂ h, list.all_inter_of_all_right _ h)
theorem subset_iff_all (s t : finset A) : s ⊆ t ↔ all s (λ x, x ∈ t) :=
iff.intro
(assume H : s ⊆ t, all_of_forall (take x, assume H1, mem_of_subset_of_mem H H1))
(assume H : all s (λ x, x ∈ t), subset_of_forall (take x, assume H1 : x ∈ s, of_mem_of_all H1 H))
(suppose s ⊆ t, all_of_forall (take x, assume H1, mem_of_subset_of_mem `s ⊆ t` H1))
(suppose H : all s (λ x, x ∈ t), subset_of_forall (take x, assume H1 : x ∈ s, of_mem_of_all H1 H))
definition decidable_subset [instance] (s t : finset A) : decidable (s ⊆ t) :=
decidable_of_decidable_of_iff _ (iff.symm !subset_iff_all)

View file

@ -112,21 +112,21 @@ assume npa, if_neg npa
theorem all_of_check_pred_eq_tt {p : A → Prop} [h : decidable_pred p] : ∀ {l : list A}, check_pred p l = tt → ∀ {a}, a ∈ l → p a
| [] eqtt a ainl := absurd ainl !not_mem_nil
| (b::l) eqtt a ainbl := by_cases
(λ pb : p b, or.elim (eq_or_mem_of_mem_cons ainbl)
(λ aeqb : a = b, by rewrite [aeqb]; exact pb)
(λ ainl : a ∈ l,
have eqtt₁ : check_pred p l = tt, by rewrite [check_pred_cons_of_pos _ pb at eqtt]; exact eqtt,
all_of_check_pred_eq_tt eqtt₁ ainl))
(λ npb : ¬ p b,
by rewrite [check_pred_cons_of_neg _ npb at eqtt]; exact (bool.no_confusion eqtt))
(suppose p b, or.elim (eq_or_mem_of_mem_cons ainbl)
(suppose a = b, by rewrite [this]; exact `p b`)
(suppose a ∈ l,
have check_pred p l = tt, by rewrite [check_pred_cons_of_pos _ `p b` at eqtt]; exact eqtt,
all_of_check_pred_eq_tt this `a ∈ l`))
(suppose ¬ p b,
by rewrite [check_pred_cons_of_neg _ this at eqtt]; exact (bool.no_confusion eqtt))
theorem ex_of_check_pred_eq_ff {p : A → Prop} [h : decidable_pred p] : ∀ {l : list A}, check_pred p l = ff → ∃ w, ¬ p w
| [] eqtt := bool.no_confusion eqtt
| (a::l) eqtt := by_cases
(λ pa : p a,
have eqtt₁ : check_pred p l = ff, by rewrite [check_pred_cons_of_pos _ pa at eqtt]; exact eqtt,
ex_of_check_pred_eq_ff eqtt₁)
(λ npa : ¬ p a, exists.intro a npa)
(suppose p a,
have check_pred p l = ff, by rewrite [check_pred_cons_of_pos _ this at eqtt]; exact eqtt,
ex_of_check_pred_eq_ff this)
(suppose ¬ p a, exists.intro a this)
end check_pred
definition decidable_forall_finite [instance] {A : Type} {p : A → Prop} [h₁ : fintype A] [h₂ : decidable_pred p]
@ -134,11 +134,11 @@ definition decidable_forall_finite [instance] {A : Type} {p : A → Prop} [h₁
match h₁ with
| fintype.mk e u c :=
match check_pred p e with
| tt := λ h : check_pred p e = tt, inl (λ a : A, all_of_check_pred_eq_tt h (c a))
| ff := λ h : check_pred p e = ff,
inr (λ n : (∀ x, p x),
obtain (a : A) (w : ¬ p a), from ex_of_check_pred_eq_ff h,
absurd (n a) w)
| tt := suppose check_pred p e = tt, inl (take a : A, all_of_check_pred_eq_tt this (c a))
| ff := suppose check_pred p e = ff,
inr (suppose ∀ x, p x,
obtain (a : A) (w : ¬ p a), from ex_of_check_pred_eq_ff `check_pred p e = ff`,
absurd (this a) w)
end rfl
end
@ -151,8 +151,8 @@ match h₁ with
obtain x px, from ex,
absurd px (all_of_check_pred_eq_tt h (c x)))
| ff := λ h : check_pred (λ a, ¬ p a) e = ff, inl (
assert aux₁ : ∃ x, ¬¬p x, from ex_of_check_pred_eq_ff h,
obtain x nnpx, from aux₁, exists.intro x (not_not_elim nnpx))
assert ∃ x, ¬¬p x, from ex_of_check_pred_eq_ff h,
obtain x nnpx, from this, exists.intro x (not_not_elim nnpx))
end rfl
end
@ -168,11 +168,11 @@ private theorem mem_of_mem_ltype_elems {A : Type} {a : A} {s : list A}
: Π {l : list A} {h : l ⊆ s} {m : a ∈ s}, mk a m ∈ ltype_elems h → a ∈ l
| [] h m lin := absurd lin !not_mem_nil
| (b::l) h m lin := or.elim (eq_or_mem_of_mem_cons lin)
(λ leq : mk a m = mk b (h b (mem_cons b l)),
as_type.no_confusion leq (λ aeqb em, by rewrite [aeqb]; exact !mem_cons))
(λ linl : mk a m ∈ ltype_elems (sub_of_cons_sub h),
have ainl : a ∈ l, from mem_of_mem_ltype_elems linl,
mem_cons_of_mem _ ainl)
(suppose mk a m = mk b (h b (mem_cons b l)),
as_type.no_confusion this (λ aeqb em, by rewrite [aeqb]; exact !mem_cons))
(suppose mk a m ∈ ltype_elems (sub_of_cons_sub h),
have a ∈ l, from mem_of_mem_ltype_elems this,
mem_cons_of_mem _ this)
private theorem nodup_ltype_elems {A : Type} {s : list A} : Π {l : list A} (d : nodup l) (h : l ⊆ s), nodup (ltype_elems h)
| [] d h := nodup_nil