mirror of
				https://github.com/nmvdw/HITs-Examples
				synced 2025-11-03 23:23:51 +01:00 
			
		
		
		
	Simplified proof of extensionalty and proofs in interface.v
This commit is contained in:
		@@ -12,7 +12,6 @@ fsets/isomorphism.v
 | 
				
			|||||||
fsets/operations.v
 | 
					fsets/operations.v
 | 
				
			||||||
fsets/operations_decidable.v
 | 
					fsets/operations_decidable.v
 | 
				
			||||||
fsets/extensionality.v
 | 
					fsets/extensionality.v
 | 
				
			||||||
fsets/extensionality_alt.v
 | 
					 | 
				
			||||||
fsets/properties.v
 | 
					fsets/properties.v
 | 
				
			||||||
fsets/properties_decidable.v
 | 
					fsets/properties_decidable.v
 | 
				
			||||||
fsets/length.v
 | 
					fsets/length.v
 | 
				
			||||||
 
 | 
				
			|||||||
@@ -6,29 +6,86 @@ Section ext.
 | 
				
			|||||||
  Context {A : Type}.
 | 
					  Context {A : Type}.
 | 
				
			||||||
  Context `{Univalence}.
 | 
					  Context `{Univalence}.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					  Lemma equiv_subset1_l (X Y : FSet A) (H1 : Y ∪ X = X) (a : A) (Ya : a ∈ Y) : a ∈ X.
 | 
				
			||||||
 | 
					  Proof.
 | 
				
			||||||
 | 
					    apply (transport (fun Z => a ∈ Z) H1 (tr(inl Ya))).
 | 
				
			||||||
 | 
					  Defined.
 | 
				
			||||||
 | 
					  
 | 
				
			||||||
 | 
					  Lemma equiv_subset1_r X : forall (Y : FSet A), (forall a, a ∈ Y -> a ∈ X) -> Y ∪ X = X.
 | 
				
			||||||
 | 
					  Proof.
 | 
				
			||||||
 | 
					    hinduction ; try (intros ; apply path_ishprop).
 | 
				
			||||||
 | 
					    - intros.
 | 
				
			||||||
 | 
					      apply nl.
 | 
				
			||||||
 | 
					    - intros b sub.
 | 
				
			||||||
 | 
					      specialize (sub b (tr idpath)).
 | 
				
			||||||
 | 
					      revert sub.
 | 
				
			||||||
 | 
					      hinduction X ; try (intros ; apply path_ishprop).
 | 
				
			||||||
 | 
					      * contradiction.
 | 
				
			||||||
 | 
					      * intros.
 | 
				
			||||||
 | 
					        strip_truncations.
 | 
				
			||||||
 | 
					        rewrite sub.
 | 
				
			||||||
 | 
					        apply union_idem.
 | 
				
			||||||
 | 
					      * intros X Y subX subY mem.
 | 
				
			||||||
 | 
					        strip_truncations.
 | 
				
			||||||
 | 
					        destruct mem as [t | t].
 | 
				
			||||||
 | 
					        ** rewrite assoc, (subX t).
 | 
				
			||||||
 | 
					           reflexivity.
 | 
				
			||||||
 | 
					        ** rewrite (comm X), assoc, (subY t).
 | 
				
			||||||
 | 
					           reflexivity.
 | 
				
			||||||
 | 
					    - intros Y1 Y2 H1 H2 H3.
 | 
				
			||||||
 | 
					      rewrite <- assoc.
 | 
				
			||||||
 | 
					      rewrite (H2 (fun a HY => H3 a (tr(inr HY)))).
 | 
				
			||||||
 | 
					      apply (H1 (fun a HY => H3 a (tr(inl HY)))).
 | 
				
			||||||
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					  Lemma eq_subset1 X Y : (Y ∪ X = X) * (X ∪ Y = Y) <~> forall (a : A), a ∈ X = a ∈ Y.
 | 
				
			||||||
 | 
					  Proof.    
 | 
				
			||||||
 | 
					    eapply equiv_iff_hprop_uncurried ; split.
 | 
				
			||||||
 | 
					    - intros [H1 H2] a.
 | 
				
			||||||
 | 
					      apply path_iff_hprop ; apply equiv_subset1_l ; assumption.
 | 
				
			||||||
 | 
					    - intros H1.
 | 
				
			||||||
 | 
					      split ; apply equiv_subset1_r ; intros.
 | 
				
			||||||
 | 
					      * rewrite H1 ; assumption.
 | 
				
			||||||
 | 
					      * rewrite <- H1 ; assumption.
 | 
				
			||||||
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					  Lemma eq_subset2 (X Y : FSet A) : X = Y <~> (Y ∪ X = X) * (X ∪ Y = Y).
 | 
				
			||||||
 | 
					  Proof.
 | 
				
			||||||
 | 
					    eapply equiv_iff_hprop_uncurried ; split.
 | 
				
			||||||
 | 
					    - intro Heq.
 | 
				
			||||||
 | 
					      split.
 | 
				
			||||||
 | 
					      * apply (ap (fun Z => Z ∪ X) Heq^ @ union_idem X).
 | 
				
			||||||
 | 
					      * apply (ap (fun Z => Z ∪ Y) Heq @ union_idem Y).
 | 
				
			||||||
 | 
					    - intros [H1 H2].
 | 
				
			||||||
 | 
					      apply (H1^ @ comm Y X @ H2).
 | 
				
			||||||
 | 
					  Defined.
 | 
				
			||||||
 | 
					  
 | 
				
			||||||
 | 
					  Theorem fset_ext (X Y : FSet A) :
 | 
				
			||||||
 | 
					    X = Y <~> forall (a : A), a ∈ X = a ∈ Y.
 | 
				
			||||||
 | 
					  Proof.
 | 
				
			||||||
 | 
					    apply (equiv_compose' (eq_subset1 X Y) (eq_subset2 X Y)).
 | 
				
			||||||
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Lemma subset_union (X Y : FSet A) :
 | 
					  Lemma subset_union (X Y : FSet A) :
 | 
				
			||||||
    X ⊆ Y -> X ∪ Y = Y.
 | 
					    X ⊆ Y -> X ∪ Y = Y.
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    hinduction X ; try (intros; apply path_forall; intro; apply set_path2).
 | 
					    hinduction X ; try (intros ; apply path_ishprop). 
 | 
				
			||||||
    - intros. apply nl.
 | 
					    - intros.
 | 
				
			||||||
 | 
					      apply nl.
 | 
				
			||||||
    - intros a.
 | 
					    - intros a.
 | 
				
			||||||
      hinduction Y ; try (intros; apply path_forall; intro; apply set_path2).
 | 
					      hinduction Y ; try (intros ; apply path_ishprop).
 | 
				
			||||||
      + intro.
 | 
					      + intro.
 | 
				
			||||||
        contradiction.
 | 
					        contradiction.
 | 
				
			||||||
      + intro a0.
 | 
					      + intros b p.
 | 
				
			||||||
        simple refine (Trunc_ind _ _).
 | 
					        strip_truncations.
 | 
				
			||||||
        intro p ; simpl.
 | 
					        rewrite p.
 | 
				
			||||||
        rewrite p; apply idem.
 | 
					        apply idem.
 | 
				
			||||||
      + intros X1 X2 IH1 IH2.
 | 
					      + intros X1 X2 IH1 IH2 t.
 | 
				
			||||||
        simple refine (Trunc_ind _ _).
 | 
					        strip_truncations.
 | 
				
			||||||
        intros [e1 | e2].
 | 
					        destruct t as [t | t].
 | 
				
			||||||
        ++ rewrite assoc.
 | 
					        ++ rewrite assoc, (IH1 t).
 | 
				
			||||||
           rewrite (IH1 e1).
 | 
					 | 
				
			||||||
           reflexivity.
 | 
					           reflexivity.
 | 
				
			||||||
        ++ rewrite comm.
 | 
					        ++ rewrite comm, <- assoc, (comm X2), (IH2 t).
 | 
				
			||||||
           rewrite <- assoc.
 | 
					 | 
				
			||||||
           rewrite (comm X2).
 | 
					 | 
				
			||||||
           rewrite (IH2 e2).
 | 
					 | 
				
			||||||
           reflexivity.
 | 
					           reflexivity.
 | 
				
			||||||
    - intros X1 X2 IH1 IH2 [G1 G2].
 | 
					    - intros X1 X2 IH1 IH2 [G1 G2].
 | 
				
			||||||
      rewrite <- assoc.
 | 
					      rewrite <- assoc.
 | 
				
			||||||
@@ -39,15 +96,16 @@ Section ext.
 | 
				
			|||||||
  Lemma subset_union_l (X : FSet A) :
 | 
					  Lemma subset_union_l (X : FSet A) :
 | 
				
			||||||
    forall Y, X ⊆ X ∪ Y.
 | 
					    forall Y, X ⊆ X ∪ Y.
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    hinduction X ; try (repeat (intro; intros; apply path_forall);
 | 
					    hinduction X ; try (intros ; apply path_ishprop).
 | 
				
			||||||
                        intro ; apply path_ishprop).
 | 
					 | 
				
			||||||
    - apply (fun _ => tt).
 | 
					    - apply (fun _ => tt).
 | 
				
			||||||
    - intros a Y.
 | 
					    - intros.
 | 
				
			||||||
      apply (tr(inl(tr idpath))).
 | 
					      apply (tr(inl(tr idpath))).
 | 
				
			||||||
    - intros X1 X2 HX1 HX2 Y.
 | 
					    - intros X1 X2 HX1 HX2 Y.
 | 
				
			||||||
      split ; unfold subset in *.
 | 
					      split ; unfold subset in *.
 | 
				
			||||||
      * rewrite <- assoc. apply HX1.
 | 
					      * rewrite <- assoc.
 | 
				
			||||||
      * rewrite (comm X1 X2). rewrite <- assoc. apply HX2.
 | 
					        apply HX1.
 | 
				
			||||||
 | 
					      * rewrite (comm X1 X2), <- assoc.
 | 
				
			||||||
 | 
					        apply HX2.
 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Lemma subset_union_equiv
 | 
					  Lemma subset_union_equiv
 | 
				
			||||||
@@ -63,90 +121,21 @@ Section ext.
 | 
				
			|||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
  
 | 
					  
 | 
				
			||||||
  Lemma subset_isIn (X Y : FSet A) :
 | 
					  Lemma subset_isIn (X Y : FSet A) :
 | 
				
			||||||
    (forall (a : A), a ∈ X -> a ∈ Y)
 | 
					    X ⊆ Y <~> forall (a : A), a ∈ X -> a ∈ Y.
 | 
				
			||||||
      <~> X ⊆ Y.
 | 
					 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    eapply equiv_iff_hprop_uncurried.
 | 
					    etransitivity.
 | 
				
			||||||
    split.
 | 
					    - apply subset_union_equiv.
 | 
				
			||||||
    - hinduction X ;
 | 
					    - eapply equiv_iff_hprop_uncurried ; split.
 | 
				
			||||||
        try (intros; repeat (apply path_forall; intro); apply path_ishprop).
 | 
					      * apply equiv_subset1_l.
 | 
				
			||||||
      + intros ; reflexivity.
 | 
					      * apply equiv_subset1_r.
 | 
				
			||||||
      + intros a f.
 | 
					 | 
				
			||||||
        apply f.
 | 
					 | 
				
			||||||
        apply tr ; reflexivity.
 | 
					 | 
				
			||||||
      + intros X1 X2 H1 H2 f.
 | 
					 | 
				
			||||||
        enough (X1 ⊆ Y).
 | 
					 | 
				
			||||||
        enough (X2 ⊆ Y).
 | 
					 | 
				
			||||||
        { split. apply X. apply X0. }
 | 
					 | 
				
			||||||
        * apply H2.
 | 
					 | 
				
			||||||
          intros a Ha.
 | 
					 | 
				
			||||||
          refine (f _ (tr _)).
 | 
					 | 
				
			||||||
          right.
 | 
					 | 
				
			||||||
          apply Ha.
 | 
					 | 
				
			||||||
        * apply H1.
 | 
					 | 
				
			||||||
          intros a Ha.
 | 
					 | 
				
			||||||
          refine (f _ (tr _)).
 | 
					 | 
				
			||||||
          left.
 | 
					 | 
				
			||||||
          apply Ha.
 | 
					 | 
				
			||||||
    - hinduction X ;
 | 
					 | 
				
			||||||
        try (intros; repeat (apply path_forall; intro); apply path_ishprop).
 | 
					 | 
				
			||||||
      + intros. contradiction.
 | 
					 | 
				
			||||||
      + intros b f a.
 | 
					 | 
				
			||||||
        simple refine (Trunc_ind _ _) ; cbn.
 | 
					 | 
				
			||||||
        intro p.
 | 
					 | 
				
			||||||
        rewrite p^ in f.
 | 
					 | 
				
			||||||
        apply f.
 | 
					 | 
				
			||||||
      + intros X1 X2 IH1 IH2 [H1 H2] a.
 | 
					 | 
				
			||||||
        simple refine (Trunc_ind _ _) ; cbn.
 | 
					 | 
				
			||||||
        intros [C1 | C2].
 | 
					 | 
				
			||||||
        ++ apply (IH1 H1 a C1).
 | 
					 | 
				
			||||||
        ++ apply (IH2 H2 a C2).
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  (** ** Extensionality proof *)
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Lemma eq_subset' (X Y : FSet A) : X = Y <~> (Y ∪ X = X) * (X ∪ Y = Y).
 | 
					 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    unshelve eapply BuildEquiv.
 | 
					 | 
				
			||||||
    { intro H'. rewrite H'. split; apply union_idem. }
 | 
					 | 
				
			||||||
    unshelve esplit.
 | 
					 | 
				
			||||||
    { intros [H1 H2]. etransitivity. apply H1^.
 | 
					 | 
				
			||||||
      rewrite comm. apply H2. }
 | 
					 | 
				
			||||||
    intro; apply path_prod; apply set_path2.
 | 
					 | 
				
			||||||
    all: intro; apply set_path2.
 | 
					 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Lemma eq_subset (X Y : FSet A) :
 | 
					  Lemma eq_subset (X Y : FSet A) :
 | 
				
			||||||
    X = Y <~> (Y ⊆ X * X ⊆ Y).
 | 
					    X = Y <~> (Y ⊆ X * X ⊆ Y).
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    transitivity ((Y ∪ X = X) * (X ∪ Y = Y)).
 | 
					    etransitivity ((Y ∪ X = X) * (X ∪ Y = Y)).
 | 
				
			||||||
    apply eq_subset'.
 | 
					    - apply eq_subset2.
 | 
				
			||||||
    symmetry.
 | 
					    - symmetry.
 | 
				
			||||||
      eapply equiv_functor_prod' ; apply subset_union_equiv.
 | 
					      eapply equiv_functor_prod' ; apply subset_union_equiv.
 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					 | 
				
			||||||
  Theorem fset_ext (X Y : FSet A) :
 | 
					 | 
				
			||||||
    X = Y <~> (forall (a : A), a ∈ X = a ∈ Y).
 | 
					 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    refine (@equiv_compose' _ _ _ _ _) ; [ | apply eq_subset ].
 | 
					 | 
				
			||||||
    refine (@equiv_compose' _ ((forall a, a ∈ Y -> a ∈ X)
 | 
					 | 
				
			||||||
                               *(forall a, a ∈ X -> a ∈ Y)) _ _ _).
 | 
					 | 
				
			||||||
    - apply equiv_iff_hprop_uncurried.
 | 
					 | 
				
			||||||
      split.
 | 
					 | 
				
			||||||
      * intros [H1 H2 a].
 | 
					 | 
				
			||||||
        specialize (H1 a) ; specialize (H2 a).
 | 
					 | 
				
			||||||
        apply path_iff_hprop.
 | 
					 | 
				
			||||||
        apply H2.
 | 
					 | 
				
			||||||
        apply H1.
 | 
					 | 
				
			||||||
      * intros H1.
 | 
					 | 
				
			||||||
        split ; intro a ; intro H2.
 | 
					 | 
				
			||||||
      + rewrite (H1 a).
 | 
					 | 
				
			||||||
        apply H2.
 | 
					 | 
				
			||||||
      + rewrite <- (H1 a).
 | 
					 | 
				
			||||||
        apply H2.
 | 
					 | 
				
			||||||
    - eapply equiv_functor_prod' ;
 | 
					 | 
				
			||||||
        apply equiv_iff_hprop_uncurried ;
 | 
					 | 
				
			||||||
        split ; apply subset_isIn.
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
End ext.
 | 
					End ext.
 | 
				
			||||||
@@ -1,68 +0,0 @@
 | 
				
			|||||||
(** Extensionality of the FSets *)
 | 
					 | 
				
			||||||
Require Import HoTT HitTactics.
 | 
					 | 
				
			||||||
Require Import representations.definition fsets.operations.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
Section ext.
 | 
					 | 
				
			||||||
  Context {A : Type}.
 | 
					 | 
				
			||||||
  Context `{Univalence}.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Lemma equiv_subset1_l (X Y : FSet A) (H1 : Y ∪ X = X) (a : A) (Ya : a ∈ Y) : a ∈ X.
 | 
					 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    apply (transport (fun Z => a ∈ Z) H1 (tr(inl Ya))).
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
  
 | 
					 | 
				
			||||||
  Lemma equiv_subset1_r X : forall (Y : FSet A), (forall a, a ∈ Y -> a ∈ X) -> Y ∪ X = X.
 | 
					 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    hinduction ; try (intros ; apply path_ishprop).
 | 
					 | 
				
			||||||
    - intros.
 | 
					 | 
				
			||||||
      apply nl.
 | 
					 | 
				
			||||||
    - intros b sub.
 | 
					 | 
				
			||||||
      specialize (sub b (tr idpath)).
 | 
					 | 
				
			||||||
      revert sub.
 | 
					 | 
				
			||||||
      hinduction X ; try (intros ; apply path_ishprop).
 | 
					 | 
				
			||||||
      * contradiction.
 | 
					 | 
				
			||||||
      * intros.
 | 
					 | 
				
			||||||
        strip_truncations.
 | 
					 | 
				
			||||||
        rewrite sub.
 | 
					 | 
				
			||||||
        apply union_idem.
 | 
					 | 
				
			||||||
      * intros X Y subX subY mem.
 | 
					 | 
				
			||||||
        strip_truncations.
 | 
					 | 
				
			||||||
        destruct mem as [t | t].
 | 
					 | 
				
			||||||
        ** rewrite assoc, (subX t).
 | 
					 | 
				
			||||||
           reflexivity.
 | 
					 | 
				
			||||||
        ** rewrite (comm X), assoc, (subY t).
 | 
					 | 
				
			||||||
           reflexivity.
 | 
					 | 
				
			||||||
    - intros Y1 Y2 H1 H2 H3.
 | 
					 | 
				
			||||||
      rewrite <- assoc.
 | 
					 | 
				
			||||||
      rewrite (H2 (fun a HY => H3 a (tr(inr HY)))).
 | 
					 | 
				
			||||||
      apply (H1 (fun a HY => H3 a (tr(inl HY)))).
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Lemma eq_subset1 X Y : (Y ∪ X = X) * (X ∪ Y = Y) <~> forall (a : A), a ∈ X = a ∈ Y.
 | 
					 | 
				
			||||||
  Proof.    
 | 
					 | 
				
			||||||
    eapply equiv_iff_hprop_uncurried ; split.
 | 
					 | 
				
			||||||
    - intros [H1 H2] a.
 | 
					 | 
				
			||||||
      apply path_iff_hprop ; apply equiv_subset1_l ; assumption.
 | 
					 | 
				
			||||||
    - intros H1.
 | 
					 | 
				
			||||||
      split ; apply equiv_subset1_r ; intros.
 | 
					 | 
				
			||||||
      * rewrite H1 ; assumption.
 | 
					 | 
				
			||||||
      * rewrite <- H1 ; assumption.
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Lemma eq_subset2 (X Y : FSet A) : X = Y <~> (Y ∪ X = X) * (X ∪ Y = Y).
 | 
					 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    eapply equiv_iff_hprop_uncurried ; split.
 | 
					 | 
				
			||||||
    - intro Heq.
 | 
					 | 
				
			||||||
      split.
 | 
					 | 
				
			||||||
      * apply (ap (fun Z => Z ∪ X) Heq^ @ union_idem X).
 | 
					 | 
				
			||||||
      * apply (ap (fun Z => Z ∪ Y) Heq @ union_idem Y).
 | 
					 | 
				
			||||||
    - intros [H1 H2].
 | 
					 | 
				
			||||||
      apply (H1^ @ comm Y X @ H2).
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
  
 | 
					 | 
				
			||||||
  Theorem fset_ext (X Y : FSet A) :
 | 
					 | 
				
			||||||
    X = Y <~> forall (a : A), a ∈ X = a ∈ Y.
 | 
					 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    apply (equiv_compose' (eq_subset1 X Y) (eq_subset2 X Y)).
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
End ext.
 | 
					 | 
				
			||||||
@@ -77,49 +77,46 @@ Section properties.
 | 
				
			|||||||
 | 
					
 | 
				
			||||||
  Ltac simplify := intros ; autounfold in * ; apply reflect_eq ; reduce.
 | 
					  Ltac simplify := intros ; autounfold in * ; apply reflect_eq ; reduce.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Definition well_defined_union : forall (A : Type) (X1 X2 Y1 Y2 : T A),
 | 
					  Definition well_defined_union (A : Type) (X1 X2 Y1 Y2 : T A) :
 | 
				
			||||||
    set_eq A X1 Y1 -> set_eq A X2 Y2 -> set_eq A (union X1 X2) (union Y1 Y2).
 | 
					    set_eq A X1 Y1 -> set_eq A X2 Y2 -> set_eq A (union X1 X2) (union Y1 Y2).
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    intros A X1 X2 Y1 Y2 HXY1 HXY2.
 | 
					    intros HXY1 HXY2.
 | 
				
			||||||
    simplify.
 | 
					    simplify.
 | 
				
			||||||
    by rewrite HXY1, HXY2.
 | 
					    by rewrite HXY1, HXY2.
 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Definition well_defined_filter : forall (A : Type) (ϕ : A -> Bool) (X Y : T A),
 | 
					  Definition well_defined_filter (A : Type) (ϕ : A -> Bool) (X Y : T A) :
 | 
				
			||||||
    set_eq A X Y -> set_eq A (filter ϕ X) (filter ϕ Y).
 | 
					    set_eq A X Y -> set_eq A (filter ϕ X) (filter ϕ Y).
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    intros A ϕ X Y HXY.
 | 
					    intros HXY.
 | 
				
			||||||
    simplify.
 | 
					    simplify.
 | 
				
			||||||
    by rewrite HXY.
 | 
					    by rewrite HXY.
 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					  Ltac reflect_equality := simplify ; eauto with lattice_hints typeclass_instances.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Lemma union_comm : forall A (X Y : T A),
 | 
					  Lemma union_comm : forall A (X Y : T A),
 | 
				
			||||||
      set_eq A (X ∪ Y) (Y ∪ X).
 | 
					      set_eq A (X ∪ Y) (Y ∪ X).
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    simplify.
 | 
					    reflect_equality.
 | 
				
			||||||
    apply comm.
 | 
					 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Lemma union_assoc : forall A (X Y Z : T A),
 | 
					  Lemma union_assoc : forall A (X Y Z : T A),
 | 
				
			||||||
    set_eq A ((X ∪ Y) ∪ Z) (X ∪ (Y ∪ Z)).
 | 
					    set_eq A ((X ∪ Y) ∪ Z) (X ∪ (Y ∪ Z)).
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    simplify.
 | 
					    reflect_equality.
 | 
				
			||||||
    symmetry.
 | 
					 | 
				
			||||||
    apply assoc.
 | 
					 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Lemma union_idem : forall A (X : T A),
 | 
					  Lemma union_idem : forall A (X : T A),
 | 
				
			||||||
    set_eq A (X ∪ X) X.
 | 
					    set_eq A (X ∪ X) X.
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    simplify.
 | 
					    reflect_equality.
 | 
				
			||||||
    apply union_idem.
 | 
					 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Lemma union_neutral : forall A (X : T A),
 | 
					  Lemma union_neutral : forall A (X : T A),
 | 
				
			||||||
    set_eq A (∅ ∪ X) X.
 | 
					    set_eq A (∅ ∪ X) X.
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    simplify.
 | 
					    reflect_equality.
 | 
				
			||||||
    apply nl.
 | 
					 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
End properties.
 | 
					End properties.
 | 
				
			||||||
@@ -258,8 +255,7 @@ Proof.
 | 
				
			|||||||
  - intros. apply path_forall; intro. apply set_path2.
 | 
					  - intros. apply path_forall; intro. apply set_path2.
 | 
				
			||||||
Defined.
 | 
					Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
Ltac buggeroff := intros;
 | 
					Ltac buggeroff := intros; apply path_ishprop.
 | 
				
			||||||
  (repeat (apply path_forall; intro)); apply set_path2.
 | 
					 | 
				
			||||||
 | 
					
 | 
				
			||||||
Instance View_max_assoc A: Associative (@max_L (View A) _).
 | 
					Instance View_max_assoc A: Associative (@max_L (View A) _).
 | 
				
			||||||
Proof.
 | 
					Proof.
 | 
				
			||||||
 
 | 
				
			|||||||
		Reference in New Issue
	
	Block a user