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:
		@@ -6,7 +6,7 @@ Section ext.
 | 
			
		||||
  Context {A : Type}.
 | 
			
		||||
  Context `{Univalence}.
 | 
			
		||||
 | 
			
		||||
  Lemma subset_union (X Y : FSet A) : 
 | 
			
		||||
  Lemma subset_union (X Y : FSet A) :
 | 
			
		||||
    X ⊆ Y -> X ∪ Y = Y.
 | 
			
		||||
  Proof.
 | 
			
		||||
    hinduction X ; try (intros; apply path_forall; intro; apply set_path2).
 | 
			
		||||
@@ -17,7 +17,7 @@ Section ext.
 | 
			
		||||
        contradiction.
 | 
			
		||||
      + intro a0.
 | 
			
		||||
        simple refine (Trunc_ind _ _).
 | 
			
		||||
        intro p ; simpl. 
 | 
			
		||||
        intro p ; simpl.
 | 
			
		||||
        rewrite p; apply idem.
 | 
			
		||||
      + intros X1 X2 IH1 IH2.
 | 
			
		||||
        simple refine (Trunc_ind _ _).
 | 
			
		||||
@@ -112,8 +112,8 @@ Section ext.
 | 
			
		||||
    unshelve esplit.
 | 
			
		||||
    { intros [H1 H2]. etransitivity. apply H1^.
 | 
			
		||||
      rewrite comm. apply H2. }
 | 
			
		||||
    intro; apply path_prod; apply set_path2. 
 | 
			
		||||
    all: intro; apply set_path2.  
 | 
			
		||||
    intro; apply path_prod; apply set_path2.
 | 
			
		||||
    all: intro; apply set_path2.
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
  Lemma eq_subset (X Y : FSet A) :
 | 
			
		||||
@@ -149,4 +149,4 @@ Section ext.
 | 
			
		||||
        split ; apply subset_isIn.
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
End ext.
 | 
			
		||||
End ext.
 | 
			
		||||
 
 | 
			
		||||
@@ -14,7 +14,7 @@ Section Iso.
 | 
			
		||||
    - apply E.
 | 
			
		||||
    - intros a x.
 | 
			
		||||
      apply ({|a|} ∪ x).
 | 
			
		||||
    - intros. cbn.  
 | 
			
		||||
    - intros. cbn.
 | 
			
		||||
      etransitivity. apply assoc.
 | 
			
		||||
      apply (ap (∪ x)).
 | 
			
		||||
      apply idem.
 | 
			
		||||
@@ -22,7 +22,7 @@ Section Iso.
 | 
			
		||||
      etransitivity. apply assoc.
 | 
			
		||||
      etransitivity. refine (ap (∪ x) _ ).
 | 
			
		||||
      apply FSet.comm.
 | 
			
		||||
      symmetry. 
 | 
			
		||||
      symmetry.
 | 
			
		||||
      apply assoc.
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
@@ -39,10 +39,10 @@ Section Iso.
 | 
			
		||||
    - apply singleton_idem.
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
  Lemma append_union: forall (x y: FSetC A), 
 | 
			
		||||
  Lemma append_union: forall (x y: FSetC A),
 | 
			
		||||
      FSetC_to_FSet (x ∪ y) = (FSetC_to_FSet x) ∪ (FSetC_to_FSet y).
 | 
			
		||||
  Proof.
 | 
			
		||||
    intros x. 
 | 
			
		||||
    intros x.
 | 
			
		||||
    hrecursion x; try (intros; apply path_forall; intro; apply set_path2).
 | 
			
		||||
    - intros. symmetry. apply nl.
 | 
			
		||||
    - intros a x HR y. unfold union, fsetc_union in *. rewrite HR. apply assoc.
 | 
			
		||||
 
 | 
			
		||||
@@ -34,7 +34,7 @@ Section operations.
 | 
			
		||||
    - apply comm.
 | 
			
		||||
    - apply nl.
 | 
			
		||||
    - apply nr.
 | 
			
		||||
    - intros; simpl. 
 | 
			
		||||
    - intros; simpl.
 | 
			
		||||
      destruct (P x).
 | 
			
		||||
      + apply idem.
 | 
			
		||||
      + apply nl.
 | 
			
		||||
@@ -61,7 +61,7 @@ Section operations.
 | 
			
		||||
    - intros ; apply nr.
 | 
			
		||||
    - intros ; apply idem.
 | 
			
		||||
  Defined.
 | 
			
		||||
        
 | 
			
		||||
 | 
			
		||||
  Definition product {A B : Type} : FSet A -> FSet B -> FSet (A * B).
 | 
			
		||||
  Proof.
 | 
			
		||||
    intros X Y.
 | 
			
		||||
@@ -76,7 +76,7 @@ Section operations.
 | 
			
		||||
    - intros ; apply nr.
 | 
			
		||||
    - intros ; apply union_idem.
 | 
			
		||||
  Defined.
 | 
			
		||||
   
 | 
			
		||||
 | 
			
		||||
  Global Instance fset_subset : forall A, hasSubset (FSet A).
 | 
			
		||||
  Proof.
 | 
			
		||||
    intros A X Y.
 | 
			
		||||
@@ -103,21 +103,6 @@ Section operations.
 | 
			
		||||
        split ; apply p.
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
  Definition map (A B : Type) (f : A -> B) : FSet A -> FSet B.
 | 
			
		||||
  Proof.
 | 
			
		||||
    hrecursion.
 | 
			
		||||
    - apply ∅.
 | 
			
		||||
    - intro a.
 | 
			
		||||
      apply {|f a|}.
 | 
			
		||||
    - apply (∪).
 | 
			
		||||
    - apply assoc.
 | 
			
		||||
    - apply comm.
 | 
			
		||||
    - apply nl.
 | 
			
		||||
    - apply nr.
 | 
			
		||||
    - intros.
 | 
			
		||||
      apply idem.
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
  Local Ltac remove_transport
 | 
			
		||||
    := intros ; apply path_forall ; intro Z ; rewrite transport_arrow
 | 
			
		||||
       ; rewrite transport_const ; simpl.
 | 
			
		||||
@@ -155,6 +140,6 @@ Section operations.
 | 
			
		||||
    - remove_transport.
 | 
			
		||||
      rewrite <- (idem (Z (x; tr 1%path))).
 | 
			
		||||
      pointwise.
 | 
			
		||||
  Defined.  
 | 
			
		||||
      
 | 
			
		||||
End operations.
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
End operations.
 | 
			
		||||
 
 | 
			
		||||
@@ -10,7 +10,7 @@ Section characterize_isIn.
 | 
			
		||||
 | 
			
		||||
  (** isIn properties *)
 | 
			
		||||
  Definition empty_isIn (a: A) : a ∈ ∅ -> Empty := idmap.
 | 
			
		||||
  
 | 
			
		||||
 | 
			
		||||
  Definition singleton_isIn (a b: A) : a ∈ {|b|} -> Trunc (-1) (a = b) := idmap.
 | 
			
		||||
 | 
			
		||||
  Definition union_isIn (X Y : FSet A) (a : A)
 | 
			
		||||
@@ -20,7 +20,7 @@ Section characterize_isIn.
 | 
			
		||||
      {|X ∪ Y & ϕ|} = {|X & ϕ|} ∪ {|Y & ϕ|}.
 | 
			
		||||
  Proof.
 | 
			
		||||
    reflexivity.
 | 
			
		||||
  Defined.  
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
  Lemma comprehension_isIn (ϕ : A -> Bool) (a : A) : forall X : FSet A,
 | 
			
		||||
      a ∈ {|X & ϕ|} = if ϕ a then a ∈ X else False_hp.
 | 
			
		||||
@@ -37,7 +37,7 @@ Section characterize_isIn.
 | 
			
		||||
        destruct c ; destruct d ; rewrite Hc, Hd ; try reflexivity
 | 
			
		||||
        ; apply path_iff_hprop ; try contradiction ; intros ; strip_truncations
 | 
			
		||||
        ; apply (false_ne_true).
 | 
			
		||||
        * apply (Hd^ @ ap ϕ X^ @ Hc). 
 | 
			
		||||
        * apply (Hd^ @ ap ϕ X^ @ Hc).
 | 
			
		||||
        * apply (Hc^ @ ap ϕ X @ Hd).
 | 
			
		||||
      }
 | 
			
		||||
      apply (X (ϕ a) (ϕ b) idpath idpath).
 | 
			
		||||
@@ -57,7 +57,7 @@ End characterize_isIn.
 | 
			
		||||
Section product_isIn.
 | 
			
		||||
  Context {A B : Type}.
 | 
			
		||||
  Context `{Univalence}.
 | 
			
		||||
  
 | 
			
		||||
 | 
			
		||||
  Lemma isIn_singleproduct (a : A) (b : B) (c : A) : forall (Y : FSet B),
 | 
			
		||||
      (a, b) ∈ (single_product c Y) = land (BuildhProp (Trunc (-1) (a = c))) (b ∈ Y).
 | 
			
		||||
  Proof.
 | 
			
		||||
@@ -65,7 +65,7 @@ Section product_isIn.
 | 
			
		||||
    - apply path_hprop ; symmetry ; apply prod_empty_r.
 | 
			
		||||
    - intros d.
 | 
			
		||||
      apply path_iff_hprop.
 | 
			
		||||
      * intros. 
 | 
			
		||||
      * intros.
 | 
			
		||||
        strip_truncations.
 | 
			
		||||
        split ; apply tr ; try (apply (ap fst X)) ; try (apply (ap snd X)).
 | 
			
		||||
      * intros [Z1 Z2].
 | 
			
		||||
@@ -93,7 +93,7 @@ Section product_isIn.
 | 
			
		||||
        ** right.
 | 
			
		||||
           split ; try (apply (tr H1)) ; try (apply Hb2).
 | 
			
		||||
  Defined.
 | 
			
		||||
  
 | 
			
		||||
 | 
			
		||||
  Definition isIn_product (a : A) (b : B) (X : FSet A) (Y : FSet B) :
 | 
			
		||||
    (a,b) ∈ (product X Y) = land (a ∈ X) (b ∈ Y).
 | 
			
		||||
  Proof.
 | 
			
		||||
@@ -147,7 +147,7 @@ Section properties.
 | 
			
		||||
  Global Instance joinsemilattice_fset : JoinSemiLattice (FSet A).
 | 
			
		||||
  Proof.
 | 
			
		||||
    split ; toHProp.
 | 
			
		||||
  Defined.  
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
  (** comprehension properties *)
 | 
			
		||||
  Lemma comprehension_false : forall (X : FSet A), {|X & fun _ => false|} = ∅.
 | 
			
		||||
@@ -176,7 +176,7 @@ Section properties.
 | 
			
		||||
  Proof.
 | 
			
		||||
    toHProp.
 | 
			
		||||
  Defined.
 | 
			
		||||
  
 | 
			
		||||
 | 
			
		||||
  Lemma merely_choice : forall X : FSet A, hor (X = ∅) (hexists (fun a => a ∈ X)).
 | 
			
		||||
  Proof.
 | 
			
		||||
    hinduction; try (intros; apply equiv_hprop_allpath ; apply _).
 | 
			
		||||
@@ -295,5 +295,5 @@ Section properties.
 | 
			
		||||
           repeat f_ap.
 | 
			
		||||
           apply path_ishprop.
 | 
			
		||||
  Defined.
 | 
			
		||||
  
 | 
			
		||||
 | 
			
		||||
End properties.
 | 
			
		||||
 
 | 
			
		||||
@@ -7,7 +7,7 @@ Section properties.
 | 
			
		||||
  Context {A : Type}.
 | 
			
		||||
 | 
			
		||||
  Definition append_nl : forall (x: FSetC A), ∅ ∪ x = x
 | 
			
		||||
    := fun _ => idpath. 
 | 
			
		||||
    := fun _ => idpath.
 | 
			
		||||
 | 
			
		||||
  Lemma append_nr : forall (x: FSetC A), x ∪ ∅ = x.
 | 
			
		||||
  Proof.
 | 
			
		||||
@@ -15,20 +15,20 @@ Section properties.
 | 
			
		||||
    -  reflexivity.
 | 
			
		||||
    -  intros. apply (ap (fun y => a;;y) X).
 | 
			
		||||
  Defined.
 | 
			
		||||
  
 | 
			
		||||
  Lemma append_assoc {H: Funext}:  
 | 
			
		||||
 | 
			
		||||
  Lemma append_assoc {H: Funext}:
 | 
			
		||||
    forall (x y z: FSetC A), x ∪ (y ∪ z) = (x ∪ y) ∪ z.
 | 
			
		||||
  Proof.
 | 
			
		||||
    hinduction
 | 
			
		||||
    ; try (intros ; apply path_forall ; intro
 | 
			
		||||
           ; apply path_forall ; intro ;  apply set_path2).
 | 
			
		||||
    - reflexivity.
 | 
			
		||||
    - intros a x HR y z. 
 | 
			
		||||
    - intros a x HR y z.
 | 
			
		||||
      specialize (HR y z).
 | 
			
		||||
      apply (ap (fun y => a;;y) HR). 
 | 
			
		||||
      apply (ap (fun y => a;;y) HR).
 | 
			
		||||
  Defined.
 | 
			
		||||
  
 | 
			
		||||
  Lemma append_singleton: forall (a: A) (x: FSetC A), 
 | 
			
		||||
 | 
			
		||||
  Lemma append_singleton: forall (a: A) (x: FSetC A),
 | 
			
		||||
      a ;; x = x ∪ (a ;; ∅).
 | 
			
		||||
  Proof.
 | 
			
		||||
    intro a. hinduction; try (intros; apply set_path2).
 | 
			
		||||
@@ -38,22 +38,22 @@ Section properties.
 | 
			
		||||
      + apply (ap (fun y => b ;; y) HR).
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
  Lemma append_comm {H: Funext}: 
 | 
			
		||||
  Lemma append_comm {H: Funext}:
 | 
			
		||||
    forall (x1 x2: FSetC A), x1 ∪ x2 = x2 ∪ x1.
 | 
			
		||||
  Proof.
 | 
			
		||||
    hinduction ;  try (intros ; apply path_forall ; intro ; apply set_path2).
 | 
			
		||||
    - intros. symmetry. apply append_nr. 
 | 
			
		||||
    - intros. symmetry. apply append_nr.
 | 
			
		||||
    - intros a x1 HR x2.
 | 
			
		||||
      etransitivity.
 | 
			
		||||
      apply (ap (fun y => a;;y) (HR x2)).  
 | 
			
		||||
      apply (ap (fun y => a;;y) (HR x2)).
 | 
			
		||||
      transitivity  ((x2 ∪ x1) ∪ (a;;∅)).
 | 
			
		||||
      + apply append_singleton. 
 | 
			
		||||
      + apply append_singleton.
 | 
			
		||||
      + etransitivity.
 | 
			
		||||
    	* symmetry. apply append_assoc.
 | 
			
		||||
    	* simple refine (ap (x2 ∪) (append_singleton _ _)^).
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
  Lemma singleton_idem: forall (a: A), 
 | 
			
		||||
  Lemma singleton_idem: forall (a: A),
 | 
			
		||||
      {|a|} ∪ {|a|} = {|a|}.
 | 
			
		||||
  Proof.
 | 
			
		||||
    unfold singleton.
 | 
			
		||||
 
 | 
			
		||||
@@ -60,7 +60,7 @@ Section operations_isIn.
 | 
			
		||||
      contradiction.
 | 
			
		||||
    - reflexivity.
 | 
			
		||||
  Defined.
 | 
			
		||||
  
 | 
			
		||||
 | 
			
		||||
  (* Union and membership *)
 | 
			
		||||
  Lemma union_isIn_b (X Y : FSet A) (a : A) :
 | 
			
		||||
    a ∈_d (X ∪ Y) = orb (a ∈_d X) (a ∈_d Y).
 | 
			
		||||
@@ -111,24 +111,24 @@ Section SetLattice.
 | 
			
		||||
    intros x y.
 | 
			
		||||
    apply (x ∪ y).
 | 
			
		||||
  Defined.
 | 
			
		||||
      
 | 
			
		||||
 | 
			
		||||
  Instance fset_min : minimum (FSet A).
 | 
			
		||||
  Proof.
 | 
			
		||||
    intros x y.
 | 
			
		||||
    apply (x ∩ y).
 | 
			
		||||
  Defined.
 | 
			
		||||
  
 | 
			
		||||
 | 
			
		||||
  Instance fset_bot : bottom (FSet A).
 | 
			
		||||
  Proof.
 | 
			
		||||
    unfold bottom.
 | 
			
		||||
    apply ∅.
 | 
			
		||||
  Defined.
 | 
			
		||||
    
 | 
			
		||||
 | 
			
		||||
  Instance lattice_fset : Lattice (FSet A).
 | 
			
		||||
  Proof.
 | 
			
		||||
    split; toBool.
 | 
			
		||||
  Defined.
 | 
			
		||||
  
 | 
			
		||||
 | 
			
		||||
End SetLattice.
 | 
			
		||||
 | 
			
		||||
(* With extensionality we can prove decidable equality *)
 | 
			
		||||
@@ -142,4 +142,4 @@ Section dec_eq.
 | 
			
		||||
    apply decidable_prod.
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
End dec_eq.
 | 
			
		||||
End dec_eq.
 | 
			
		||||
 
 | 
			
		||||
		Reference in New Issue
	
	Block a user