mirror of
				https://github.com/nmvdw/HITs-Examples
				synced 2025-11-04 07:33:51 +01:00 
			
		
		
		
	Clean up trailing whitespaces and an unused definition.
This commit is contained in:
		@@ -7,7 +7,7 @@ Module Export T.
 | 
			
		||||
 | 
			
		||||
    Private Inductive T (B : Type) : Type :=
 | 
			
		||||
    | b : T B
 | 
			
		||||
    | c : T B.    
 | 
			
		||||
    | c : T B.
 | 
			
		||||
 | 
			
		||||
    Axiom p : A -> b A = c A.
 | 
			
		||||
    Axiom setT : IsHSet (T A).
 | 
			
		||||
@@ -23,7 +23,7 @@ Module Export T.
 | 
			
		||||
    Variable (bP : P (b A)).
 | 
			
		||||
    Variable (cP : P (c A)).
 | 
			
		||||
    Variable (pP : forall a : A, (p a) # bP = cP).
 | 
			
		||||
    
 | 
			
		||||
 | 
			
		||||
    (* Induction principle *)
 | 
			
		||||
    Fixpoint T_ind
 | 
			
		||||
             (x : T A)
 | 
			
		||||
@@ -31,7 +31,7 @@ Module Export T.
 | 
			
		||||
      : P x
 | 
			
		||||
      := (match x return _ -> _ -> P x with
 | 
			
		||||
          | b => fun _ _ => bP
 | 
			
		||||
          | c => fun _ _ => cP                              
 | 
			
		||||
          | c => fun _ _ => cP
 | 
			
		||||
          end) pP H.
 | 
			
		||||
 | 
			
		||||
    Axiom T_ind_beta_p : forall (a : A),
 | 
			
		||||
@@ -68,7 +68,7 @@ Module Export T.
 | 
			
		||||
  End T_recursion.
 | 
			
		||||
 | 
			
		||||
  Instance T_recursion A : HitRecursion (T A)
 | 
			
		||||
    := {indTy := _; recTy := _; 
 | 
			
		||||
    := {indTy := _; recTy := _;
 | 
			
		||||
        H_inductor := T_ind A; H_recursor := T_rec A }.
 | 
			
		||||
 | 
			
		||||
End T.
 | 
			
		||||
@@ -119,44 +119,44 @@ Section merely_dec_lem.
 | 
			
		||||
 | 
			
		||||
  Local Ltac f_prop := apply path_forall ; intro ; apply path_ishprop.
 | 
			
		||||
 | 
			
		||||
  Lemma transport_code_b_x (a : A) : 
 | 
			
		||||
  Lemma transport_code_b_x (a : A) :
 | 
			
		||||
    transport code_b (p a) = fun _ => a.
 | 
			
		||||
  Proof.
 | 
			
		||||
    f_prop.
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
  Lemma transport_code_c_x (a : A) : 
 | 
			
		||||
  Lemma transport_code_c_x (a : A) :
 | 
			
		||||
    transport code_c (p a) = fun _ => tt.
 | 
			
		||||
  Proof.
 | 
			
		||||
    f_prop.    
 | 
			
		||||
    f_prop.
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
  Lemma transport_code_c_x_V (a : A) : 
 | 
			
		||||
  Lemma transport_code_c_x_V (a : A) :
 | 
			
		||||
    transport code_c (p a)^ = fun _ => a.
 | 
			
		||||
  Proof. 
 | 
			
		||||
    f_prop.    
 | 
			
		||||
  Proof.
 | 
			
		||||
    f_prop.
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
  Lemma transport_code_x_b (a : A) : 
 | 
			
		||||
  Lemma transport_code_x_b (a : A) :
 | 
			
		||||
    transport (fun x => code x (b A)) (p a) = fun _ => a.
 | 
			
		||||
  Proof.
 | 
			
		||||
    f_prop.
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
  Lemma transport_code_x_c (a : A) : 
 | 
			
		||||
  Lemma transport_code_x_c (a : A) :
 | 
			
		||||
    transport (fun x => code x (c A)) (p a) = fun _ => tt.
 | 
			
		||||
  Proof.
 | 
			
		||||
    f_prop.
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
  Lemma transport_code_x_c_V (a : A) : 
 | 
			
		||||
  Lemma transport_code_x_c_V (a : A) :
 | 
			
		||||
    transport (fun x => code x (c A)) (p a)^ = fun _ => a.
 | 
			
		||||
  Proof.
 | 
			
		||||
    f_prop.
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
  Lemma ap_diag {B : Type} {x y : B} (p : x = y) :
 | 
			
		||||
    ap (fun x : B => (x, x)) p = path_prod' p p.   
 | 
			
		||||
    ap (fun x : B => (x, x)) p = path_prod' p p.
 | 
			
		||||
  Proof.
 | 
			
		||||
      by path_induction.
 | 
			
		||||
  Defined.
 | 
			
		||||
@@ -217,7 +217,7 @@ Section merely_dec_lem.
 | 
			
		||||
      refine (transport_arrow _ _ _ @ _).
 | 
			
		||||
      refine (transport_paths_FlFr _ _ @ _).
 | 
			
		||||
      rewrite transport_code_c_x_V.
 | 
			
		||||
      hott_simpl.      
 | 
			
		||||
      hott_simpl.
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
  Lemma transport_paths_FlFr_trunc :
 | 
			
		||||
@@ -229,7 +229,7 @@ Section merely_dec_lem.
 | 
			
		||||
    refine (ap tr _).
 | 
			
		||||
    exact ((concat_1p r)^ @ (concat_p1 (1 @ r))^).
 | 
			
		||||
  Defined.
 | 
			
		||||
  
 | 
			
		||||
 | 
			
		||||
  Definition decode : forall (x y : T A), code x y -> x = y.
 | 
			
		||||
  Proof.
 | 
			
		||||
    simple refine (T_ind _ _ _ _ _ _); simpl.
 | 
			
		||||
@@ -248,7 +248,7 @@ Section merely_dec_lem.
 | 
			
		||||
        f_ap.
 | 
			
		||||
        refine (ap (fun x => (p x)) _).
 | 
			
		||||
        apply path_ishprop.
 | 
			
		||||
      + intro.        
 | 
			
		||||
      + intro.
 | 
			
		||||
        rewrite transport_code_x_c_V.
 | 
			
		||||
        hott_simpl.
 | 
			
		||||
      + intro b.
 | 
			
		||||
@@ -264,7 +264,7 @@ Section merely_dec_lem.
 | 
			
		||||
    intros p. induction p.
 | 
			
		||||
    simpl. revert u. simple refine (T_ind _ _ _ _ _ _).
 | 
			
		||||
    + simpl. reflexivity.
 | 
			
		||||
    + simpl. reflexivity.    
 | 
			
		||||
    + simpl. reflexivity.
 | 
			
		||||
    + intro a.
 | 
			
		||||
      apply set_path2.
 | 
			
		||||
  Defined.
 | 
			
		||||
@@ -278,12 +278,12 @@ Section merely_dec_lem.
 | 
			
		||||
      + simpl. intro a. apply path_ishprop.
 | 
			
		||||
      + intro a. apply path_forall; intros ?. apply set_path2.
 | 
			
		||||
    - simple refine (T_ind _ _ _ _ _ _).
 | 
			
		||||
      + simpl. intro a. apply path_ishprop. 
 | 
			
		||||
      + simpl. apply path_ishprop. 
 | 
			
		||||
      + simpl. intro a. apply path_ishprop.
 | 
			
		||||
      + simpl. apply path_ishprop.
 | 
			
		||||
      + intro a. apply path_forall; intros ?. apply set_path2.
 | 
			
		||||
    - intro a. repeat (apply path_forall; intros ?). apply set_path2.
 | 
			
		||||
  Defined.
 | 
			
		||||
  
 | 
			
		||||
 | 
			
		||||
 | 
			
		||||
  Instance T_hprop (a : A) : IsHProp (b A = c A).
 | 
			
		||||
  Proof.
 | 
			
		||||
@@ -307,7 +307,7 @@ Section merely_dec_lem.
 | 
			
		||||
    rewrite ?decode_encode in H1.
 | 
			
		||||
    apply H1.
 | 
			
		||||
  Defined.
 | 
			
		||||
  
 | 
			
		||||
 | 
			
		||||
  Lemma equiv_pathspace_T : (b A = c A) = A.
 | 
			
		||||
  Proof.
 | 
			
		||||
    apply path_iff_ishprop.
 | 
			
		||||
 
 | 
			
		||||
@@ -3,7 +3,7 @@ Require Import HoTT HitTactics.
 | 
			
		||||
Require Export notation.
 | 
			
		||||
 | 
			
		||||
Module Export FSetC.
 | 
			
		||||
  
 | 
			
		||||
 | 
			
		||||
  Section FSetC.
 | 
			
		||||
    Private Inductive FSetC (A : Type) : Type :=
 | 
			
		||||
    | Nil : FSetC A
 | 
			
		||||
@@ -14,9 +14,9 @@ Module Export FSetC.
 | 
			
		||||
    Variable A : Type.
 | 
			
		||||
    Arguments Cns {_} _ _.
 | 
			
		||||
    Infix ";;" := Cns (at level 8, right associativity).
 | 
			
		||||
    
 | 
			
		||||
 | 
			
		||||
    Axiom dupl : forall (a : A) (x : FSetC A),
 | 
			
		||||
        a ;; a ;; x = a ;; x. 
 | 
			
		||||
        a ;; a ;; x = a ;; x.
 | 
			
		||||
 | 
			
		||||
    Axiom comm : forall (a b : A) (x : FSetC A),
 | 
			
		||||
        a ;; b ;; x = b ;; a ;; x.
 | 
			
		||||
@@ -41,9 +41,9 @@ Module Export FSetC.
 | 
			
		||||
    Variable (duplP : forall (a: A) (x: FSetC A) (px : P x),
 | 
			
		||||
	         dupl a x # cnsP a (a;;x) (cnsP a x px) = cnsP a x px).
 | 
			
		||||
    Variable (commP : forall (a b: A) (x: FSetC A) (px: P x),
 | 
			
		||||
		 comm a b x # cnsP a (b;;x) (cnsP b x px) = 
 | 
			
		||||
		 comm a b x # cnsP a (b;;x) (cnsP b x px) =
 | 
			
		||||
		 cnsP b (a;;x) (cnsP a x px)).
 | 
			
		||||
    
 | 
			
		||||
 | 
			
		||||
    (* Induction principle *)
 | 
			
		||||
    Fixpoint FSetC_ind
 | 
			
		||||
             (x : FSetC A)
 | 
			
		||||
@@ -76,18 +76,18 @@ Module Export FSetC.
 | 
			
		||||
    (* Recursion principle *)
 | 
			
		||||
    Definition FSetC_rec : FSetC A -> P.
 | 
			
		||||
    Proof.
 | 
			
		||||
      simple refine (FSetC_ind _ _ _ _ _  _ _ ); 
 | 
			
		||||
      simple refine (FSetC_ind _ _ _ _ _  _ _ );
 | 
			
		||||
        try (intros; simple refine ((transport_const _ _) @ _ ));  cbn.
 | 
			
		||||
      - apply nil.
 | 
			
		||||
      - apply 	(fun a => fun _ => cns a). 
 | 
			
		||||
      - apply 	(fun a => fun _ => cns a).
 | 
			
		||||
      - apply duplP.
 | 
			
		||||
      - apply commP. 
 | 
			
		||||
      - apply commP.
 | 
			
		||||
    Defined.
 | 
			
		||||
 | 
			
		||||
    Definition FSetC_rec_beta_dupl : forall (a: A) (x : FSetC A),
 | 
			
		||||
        ap FSetC_rec (dupl a x) = duplP a (FSetC_rec x).
 | 
			
		||||
    Proof.
 | 
			
		||||
      intros. 
 | 
			
		||||
      intros.
 | 
			
		||||
      eapply (cancelL (transport_const (dupl a x) _)).
 | 
			
		||||
      simple refine ((apD_const _ _)^ @ _).
 | 
			
		||||
      apply FSetC_ind_beta_dupl.
 | 
			
		||||
@@ -96,7 +96,7 @@ Module Export FSetC.
 | 
			
		||||
    Definition FSetC_rec_beta_comm : forall (a b: A) (x : FSetC A),
 | 
			
		||||
        ap FSetC_rec (comm a b x) = commP a b (FSetC_rec x).
 | 
			
		||||
    Proof.
 | 
			
		||||
      intros. 
 | 
			
		||||
      intros.
 | 
			
		||||
      eapply (cancelL (transport_const (comm a b x) _)).
 | 
			
		||||
      simple refine ((apD_const _ _)^ @ _).
 | 
			
		||||
      apply FSetC_ind_beta_comm.
 | 
			
		||||
@@ -107,7 +107,7 @@ Module Export FSetC.
 | 
			
		||||
 | 
			
		||||
  Instance FSetC_recursion A : HitRecursion (FSetC A) :=
 | 
			
		||||
    {
 | 
			
		||||
      indTy := _; recTy := _; 
 | 
			
		||||
      indTy := _; recTy := _;
 | 
			
		||||
      H_inductor := FSetC_ind A; H_recursor := FSetC_rec A
 | 
			
		||||
    }.
 | 
			
		||||
 | 
			
		||||
 
 | 
			
		||||
@@ -14,7 +14,7 @@ Module Export FSet.
 | 
			
		||||
    Global Instance fset_union : forall A, hasUnion (FSet A) := U.
 | 
			
		||||
 | 
			
		||||
    Variable A : Type.
 | 
			
		||||
    
 | 
			
		||||
 | 
			
		||||
    Axiom assoc : forall (x y z : FSet A),
 | 
			
		||||
        x ∪ (y ∪ z) = (x ∪ y) ∪ z.
 | 
			
		||||
 | 
			
		||||
@@ -38,7 +38,7 @@ Module Export FSet.
 | 
			
		||||
  Arguments comm {_} _ _.
 | 
			
		||||
  Arguments nl {_} _.
 | 
			
		||||
  Arguments nr {_} _.
 | 
			
		||||
  Arguments idem {_} _.  
 | 
			
		||||
  Arguments idem {_} _.
 | 
			
		||||
 | 
			
		||||
  Section FSet_induction.
 | 
			
		||||
    Variable A: Type.
 | 
			
		||||
@@ -47,22 +47,22 @@ Module Export FSet.
 | 
			
		||||
    Variable  (eP : P ∅).
 | 
			
		||||
    Variable  (lP : forall a: A, P {|a|}).
 | 
			
		||||
    Variable  (uP : forall (x y: FSet A), P x -> P y -> P (x ∪ y)).
 | 
			
		||||
    Variable  (assocP : forall (x y z : FSet A) 
 | 
			
		||||
    Variable  (assocP : forall (x y z : FSet A)
 | 
			
		||||
                               (px: P x) (py: P y) (pz: P z),
 | 
			
		||||
                  assoc x y z #
 | 
			
		||||
                        (uP x (y ∪ z) px (uP y z py pz)) 
 | 
			
		||||
                  = 
 | 
			
		||||
                        (uP x (y ∪ z) px (uP y z py pz))
 | 
			
		||||
                  =
 | 
			
		||||
                  (uP (x ∪ y) z (uP x y px py) pz)).
 | 
			
		||||
    Variable  (commP : forall (x y: FSet A) (px: P x) (py: P y),
 | 
			
		||||
                  comm x y #
 | 
			
		||||
                       uP x y px py = uP y x py px).
 | 
			
		||||
    Variable  (nlP : forall (x : FSet A) (px: P x), 
 | 
			
		||||
    Variable  (nlP : forall (x : FSet A) (px: P x),
 | 
			
		||||
                  nl x # uP ∅ x eP px = px).
 | 
			
		||||
    Variable  (nrP : forall (x : FSet A) (px: P x), 
 | 
			
		||||
    Variable  (nrP : forall (x : FSet A) (px: P x),
 | 
			
		||||
                  nr x # uP x ∅ px eP = px).
 | 
			
		||||
    Variable  (idemP : forall (x : A), 
 | 
			
		||||
    Variable  (idemP : forall (x : A),
 | 
			
		||||
                  idem x # uP {|x|} {|x|} (lP x) (lP x) = lP x).
 | 
			
		||||
    
 | 
			
		||||
 | 
			
		||||
    (* Induction principle *)
 | 
			
		||||
    Fixpoint FSet_ind
 | 
			
		||||
             (x : FSet A)
 | 
			
		||||
@@ -119,7 +119,7 @@ Module Export FSet.
 | 
			
		||||
    Defined.
 | 
			
		||||
 | 
			
		||||
    Definition FSet_rec_beta_assoc : forall (x y z : FSet A),
 | 
			
		||||
        ap FSet_rec (assoc x y z) 
 | 
			
		||||
        ap FSet_rec (assoc x y z)
 | 
			
		||||
        =
 | 
			
		||||
        assocP (FSet_rec x) (FSet_rec y) (FSet_rec z).
 | 
			
		||||
    Proof.
 | 
			
		||||
@@ -131,7 +131,7 @@ Module Export FSet.
 | 
			
		||||
    Defined.
 | 
			
		||||
 | 
			
		||||
    Definition FSet_rec_beta_comm : forall (x y : FSet A),
 | 
			
		||||
        ap FSet_rec (comm x y) 
 | 
			
		||||
        ap FSet_rec (comm x y)
 | 
			
		||||
        =
 | 
			
		||||
        commP (FSet_rec x) (FSet_rec y).
 | 
			
		||||
    Proof.
 | 
			
		||||
@@ -143,7 +143,7 @@ Module Export FSet.
 | 
			
		||||
    Defined.
 | 
			
		||||
 | 
			
		||||
    Definition FSet_rec_beta_nl : forall (x : FSet A),
 | 
			
		||||
        ap FSet_rec (nl x) 
 | 
			
		||||
        ap FSet_rec (nl x)
 | 
			
		||||
        =
 | 
			
		||||
        nlP (FSet_rec x).
 | 
			
		||||
    Proof.
 | 
			
		||||
@@ -155,7 +155,7 @@ Module Export FSet.
 | 
			
		||||
    Defined.
 | 
			
		||||
 | 
			
		||||
    Definition FSet_rec_beta_nr : forall (x : FSet A),
 | 
			
		||||
        ap FSet_rec (nr x) 
 | 
			
		||||
        ap FSet_rec (nr x)
 | 
			
		||||
        =
 | 
			
		||||
        nrP (FSet_rec x).
 | 
			
		||||
    Proof.
 | 
			
		||||
@@ -167,7 +167,7 @@ Module Export FSet.
 | 
			
		||||
    Defined.
 | 
			
		||||
 | 
			
		||||
    Definition FSet_rec_beta_idem : forall (a : A),
 | 
			
		||||
        ap FSet_rec (idem a) 
 | 
			
		||||
        ap FSet_rec (idem a)
 | 
			
		||||
        =
 | 
			
		||||
        idemP a.
 | 
			
		||||
    Proof.
 | 
			
		||||
@@ -177,12 +177,12 @@ Module Export FSet.
 | 
			
		||||
      simple refine ((apD_const _ _)^ @ _).
 | 
			
		||||
      apply FSet_ind_beta_idem.
 | 
			
		||||
    Defined.
 | 
			
		||||
    
 | 
			
		||||
 | 
			
		||||
  End FSet_recursion.
 | 
			
		||||
 | 
			
		||||
  Instance FSet_recursion A : HitRecursion (FSet A) :=
 | 
			
		||||
    {
 | 
			
		||||
      indTy := _; recTy := _; 
 | 
			
		||||
      indTy := _; recTy := _;
 | 
			
		||||
      H_inductor := FSet_ind A; H_recursor := FSet_rec A
 | 
			
		||||
    }.
 | 
			
		||||
 | 
			
		||||
@@ -200,5 +200,5 @@ Proof.
 | 
			
		||||
    rewrite P.
 | 
			
		||||
    rewrite (comm y x).
 | 
			
		||||
    rewrite <- (assoc x y y).
 | 
			
		||||
    f_ap. 
 | 
			
		||||
Defined.
 | 
			
		||||
    f_ap.
 | 
			
		||||
Defined.
 | 
			
		||||
 
 | 
			
		||||
		Reference in New Issue
	
	Block a user