mirror of
				https://github.com/nmvdw/HITs-Examples
				synced 2025-11-04 07:33:51 +01:00 
			
		
		
		
	Simplified proof of extensionality
This commit is contained in:
		@@ -13,64 +13,41 @@ Section lor_props.
 | 
				
			|||||||
  Context `{Univalence}.
 | 
					  Context `{Univalence}.
 | 
				
			||||||
  Variable X Y Z : hProp.
 | 
					  Variable X Y Z : hProp.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					  Local Ltac lor_intros :=
 | 
				
			||||||
 | 
					    let x := fresh in intro x
 | 
				
			||||||
 | 
					             ; repeat (strip_truncations ; destruct x as [x | x]).
 | 
				
			||||||
 | 
					  
 | 
				
			||||||
 | 
					  
 | 
				
			||||||
  Lemma lor_assoc : (X ∨ Y) ∨ Z = X ∨ Y ∨ Z.
 | 
					  Lemma lor_assoc : (X ∨ Y) ∨ Z = X ∨ Y ∨ Z.
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    apply path_iff_hprop ; cbn.
 | 
					    apply path_iff_hprop ; lor_intros
 | 
				
			||||||
    * simple refine (Trunc_ind _ _).
 | 
					    ; apply tr ; auto
 | 
				
			||||||
      intros [xy | z] ; cbn.
 | 
					    ; try (left ; apply tr)
 | 
				
			||||||
      + simple refine (Trunc_ind _ _ xy).
 | 
					    ; try (right ; apply tr) ; auto.
 | 
				
			||||||
        intros [x | y].
 | 
					 | 
				
			||||||
        ++ apply (tr (inl x)).
 | 
					 | 
				
			||||||
        ++ apply (tr (inr (tr (inl y)))).
 | 
					 | 
				
			||||||
      + apply (tr (inr (tr (inr z)))).
 | 
					 | 
				
			||||||
    * simple refine (Trunc_ind _ _).
 | 
					 | 
				
			||||||
      intros [x | yz] ; cbn.
 | 
					 | 
				
			||||||
      + apply (tr (inl (tr (inl x)))).
 | 
					 | 
				
			||||||
      + simple refine (Trunc_ind _ _ yz).
 | 
					 | 
				
			||||||
        intros [y | z].
 | 
					 | 
				
			||||||
        ++ apply (tr (inl (tr (inr y)))).
 | 
					 | 
				
			||||||
        ++ apply (tr (inr z)).
 | 
					 | 
				
			||||||
  Defined.  
 | 
					  Defined.  
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Lemma lor_comm : X ∨ Y = Y ∨ X.
 | 
					  Lemma lor_comm : X ∨ Y = Y ∨ X.
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    apply path_iff_hprop ; cbn.
 | 
					    apply path_iff_hprop ; lor_intros
 | 
				
			||||||
    * simple refine (Trunc_ind _ _).
 | 
					    ; apply tr ; auto.
 | 
				
			||||||
      intros [x | y].
 | 
					 | 
				
			||||||
      + apply (tr (inr x)).
 | 
					 | 
				
			||||||
      + apply (tr (inl y)).
 | 
					 | 
				
			||||||
    * simple refine (Trunc_ind _ _).
 | 
					 | 
				
			||||||
      intros [y | x].
 | 
					 | 
				
			||||||
      + apply (tr (inr y)).
 | 
					 | 
				
			||||||
      + apply (tr (inl x)).
 | 
					 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Lemma lor_nl : False_hp ∨ X = X.
 | 
					  Lemma lor_nl : False_hp ∨ X = X.
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    apply path_iff_hprop ; cbn.
 | 
					    apply path_iff_hprop ; lor_intros ; try contradiction
 | 
				
			||||||
    * simple refine (Trunc_ind _ _).
 | 
					    ; try (refine (tr(inr _))) ; auto.
 | 
				
			||||||
      intros [ | x].
 | 
					 | 
				
			||||||
      + apply Empty_rec.
 | 
					 | 
				
			||||||
      + apply x.
 | 
					 | 
				
			||||||
    * apply (fun x => tr (inr x)).
 | 
					 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Lemma lor_nr : X ∨ False_hp = X.
 | 
					  Lemma lor_nr : X ∨ False_hp = X.
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    apply path_iff_hprop ; cbn.
 | 
					    apply path_iff_hprop ; lor_intros ; try contradiction
 | 
				
			||||||
    * simple refine (Trunc_ind _ _).
 | 
					    ; try (refine (tr(inl _))) ; auto.
 | 
				
			||||||
      intros [x | ].
 | 
					 | 
				
			||||||
      + apply x.
 | 
					 | 
				
			||||||
      + apply Empty_rec.
 | 
					 | 
				
			||||||
    * apply (fun x => tr (inl x)).
 | 
					 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Lemma lor_idem : X ∨ X = X.
 | 
					  Lemma lor_idem : X ∨ X = X.
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    apply path_iff_hprop ; cbn.
 | 
					    apply path_iff_hprop ; lor_intros
 | 
				
			||||||
    - simple refine (Trunc_ind _ _).
 | 
					    ; try(refine (tr(inl _))) ; auto.
 | 
				
			||||||
      intros [x | x] ; apply x.
 | 
					 | 
				
			||||||
    - apply (fun x => tr (inl x)).
 | 
					 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
End lor_props.
 | 
					End lor_props.
 | 
				
			||||||
 
 | 
				
			|||||||
@@ -6,23 +6,20 @@ Section ext.
 | 
				
			|||||||
  Context {A : Type}.
 | 
					  Context {A : Type}.
 | 
				
			||||||
  Context `{Univalence}.
 | 
					  Context `{Univalence}.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Lemma equiv_subset_l : forall (X Y : FSet A), Y ∪ X = X -> forall a, a ∈ Y -> a ∈ X.
 | 
					  Lemma equiv_subset1_l (X Y : FSet A) (H1 : Y ∪ X = X) (a : A) (Ya : a ∈ Y) : a ∈ X.
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    intros X Y H1 a Ya.
 | 
					    apply (transport (fun Z => a ∈ Z) H1 (tr(inl Ya))).
 | 
				
			||||||
    rewrite <- H1.
 | 
					 | 
				
			||||||
    apply (tr(inl Ya)).
 | 
					 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
  
 | 
					  
 | 
				
			||||||
  Lemma equiv_subset_r : forall (X Y : FSet A), (forall a, a ∈ Y -> a ∈ X) -> Y ∪ X = X.
 | 
					  Lemma equiv_subset1_r X : forall (Y : FSet A), (forall a, a ∈ Y -> a ∈ X) -> Y ∪ X = X.
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    intros X.
 | 
					    hinduction ; try (intros ; apply path_ishprop).
 | 
				
			||||||
    hinduction ; try (intros ; apply path_forall ; intro ; apply path_ishprop).
 | 
					 | 
				
			||||||
    - intros.
 | 
					    - intros.
 | 
				
			||||||
      apply nl.
 | 
					      apply nl.
 | 
				
			||||||
    - intros b sub.
 | 
					    - intros b sub.
 | 
				
			||||||
      specialize (sub b (tr idpath)).
 | 
					      specialize (sub b (tr idpath)).
 | 
				
			||||||
      revert sub.
 | 
					      revert sub.
 | 
				
			||||||
      hinduction X ; try (intros ; apply path_forall ; intro ; apply path_ishprop).
 | 
					      hinduction X ; try (intros ; apply path_ishprop).
 | 
				
			||||||
      * contradiction.
 | 
					      * contradiction.
 | 
				
			||||||
      * intros.
 | 
					      * intros.
 | 
				
			||||||
        strip_truncations.
 | 
					        strip_truncations.
 | 
				
			||||||
@@ -30,45 +27,42 @@ Section ext.
 | 
				
			|||||||
        apply union_idem.
 | 
					        apply union_idem.
 | 
				
			||||||
      * intros X Y subX subY mem.
 | 
					      * intros X Y subX subY mem.
 | 
				
			||||||
        strip_truncations.
 | 
					        strip_truncations.
 | 
				
			||||||
        destruct mem.
 | 
					        destruct mem as [t | t].
 | 
				
			||||||
        ** rewrite assoc.
 | 
					        ** rewrite assoc, (subX t).
 | 
				
			||||||
           rewrite (subX t).
 | 
					 | 
				
			||||||
           reflexivity.
 | 
					           reflexivity.
 | 
				
			||||||
        ** rewrite (comm X).
 | 
					        ** rewrite (comm X), assoc, (subY t).
 | 
				
			||||||
           rewrite assoc.
 | 
					 | 
				
			||||||
           rewrite (subY t).
 | 
					 | 
				
			||||||
           reflexivity.
 | 
					           reflexivity.
 | 
				
			||||||
    - intros Y1 Y2 H1 H2 H3.
 | 
					    - intros Y1 Y2 H1 H2 H3.
 | 
				
			||||||
      rewrite <- assoc.
 | 
					      rewrite <- assoc.
 | 
				
			||||||
      rewrite (H2 (fun a HY => H3 a (tr(inr HY)))).
 | 
					      rewrite (H2 (fun a HY => H3 a (tr(inr HY)))).
 | 
				
			||||||
      rewrite (H1 (fun a HY => H3 a (tr(inl HY)))).
 | 
					      apply (H1 (fun a HY => H3 a (tr(inl HY)))).
 | 
				
			||||||
      reflexivity.
 | 
					 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Lemma eq_subset' (X Y : FSet A) : X = Y <~> (Y ∪ X = X) * (X ∪ Y = Y).
 | 
					  Lemma eq_subset1 X Y : (Y ∪ X = X) * (X ∪ Y = Y) <~> forall (a : A), a ∈ X = a ∈ Y.
 | 
				
			||||||
  Proof.    
 | 
					  Proof.    
 | 
				
			||||||
    unshelve eapply BuildEquiv.
 | 
					    eapply equiv_iff_hprop_uncurried ; split.
 | 
				
			||||||
    { intro H'. rewrite H'. split; apply union_idem. }
 | 
					    - intros [H1 H2] a.
 | 
				
			||||||
    unshelve esplit.
 | 
					      apply path_iff_hprop ; apply equiv_subset1_l ; assumption.
 | 
				
			||||||
    { intros [H1 H2]. etransitivity. apply H1^.
 | 
					    - intros H1.
 | 
				
			||||||
      rewrite comm. apply H2. }
 | 
					      split ; apply equiv_subset1_r ; intros.
 | 
				
			||||||
    intro; apply path_prod; apply set_path2.
 | 
					      * rewrite H1 ; assumption.
 | 
				
			||||||
    all: intro; apply set_path2.
 | 
					      * 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.
 | 
					  Defined.
 | 
				
			||||||
  
 | 
					  
 | 
				
			||||||
  Theorem fset_ext (X Y : FSet A) :
 | 
					  Theorem fset_ext (X Y : FSet A) :
 | 
				
			||||||
    X = Y <~> (forall (a : A), a ∈ X = a ∈ Y).
 | 
					    X = Y <~> forall (a : A), a ∈ X = a ∈ Y.
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    refine (@equiv_compose' _ _ _ _ _) ; [ | apply eq_subset' ].
 | 
					    apply (equiv_compose' (eq_subset1 X Y) (eq_subset2 X Y)).
 | 
				
			||||||
    eapply equiv_iff_hprop_uncurried ; split.
 | 
					 | 
				
			||||||
    - intros [H1 H2 a].
 | 
					 | 
				
			||||||
      apply path_iff_hprop ; apply equiv_subset_l ; assumption.
 | 
					 | 
				
			||||||
    - intros H1.
 | 
					 | 
				
			||||||
      split ; apply equiv_subset_r.
 | 
					 | 
				
			||||||
      * intros a Ya.
 | 
					 | 
				
			||||||
        rewrite (H1 a) ; assumption.
 | 
					 | 
				
			||||||
      * intros a Xa.
 | 
					 | 
				
			||||||
        rewrite <- (H1 a) ; assumption.
 | 
					 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					 | 
				
			||||||
End ext.
 | 
					End ext.
 | 
				
			||||||
		Reference in New Issue
	
	Block a user