mirror of
				https://github.com/nmvdw/HITs-Examples
				synced 2025-11-03 23:23:51 +01:00 
			
		
		
		
	Improved notatio
This commit is contained in:
		@@ -1,5 +1,6 @@
 | 
				
			|||||||
-R . "" COQC = hoqc COQDEP = hoqdep
 | 
					-R . "" COQC = hoqc COQDEP = hoqdep
 | 
				
			||||||
-R ../prelude ""
 | 
					-R ../prelude ""
 | 
				
			||||||
 | 
					notation.v
 | 
				
			||||||
lattice.v
 | 
					lattice.v
 | 
				
			||||||
disjunction.v
 | 
					disjunction.v
 | 
				
			||||||
representations/bad.v
 | 
					representations/bad.v
 | 
				
			||||||
 
 | 
				
			|||||||
@@ -1,6 +1,6 @@
 | 
				
			|||||||
(* Logical disjunction in HoTT (see ch. 3 of the book) *)
 | 
					(* Logical disjunction in HoTT (see ch. 3 of the book) *)
 | 
				
			||||||
Require Import HoTT.
 | 
					Require Import HoTT.
 | 
				
			||||||
Require Import lattice.
 | 
					Require Import lattice notation.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
Instance lor : maximum hProp := fun X Y => BuildhProp (Trunc (-1) (sum X Y)).
 | 
					Instance lor : maximum hProp := fun X Y => BuildhProp (Trunc (-1) (sum X Y)).
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 
 | 
				
			|||||||
@@ -12,25 +12,35 @@ Section Iso.
 | 
				
			|||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    hrecursion.
 | 
					    hrecursion.
 | 
				
			||||||
    - apply E.
 | 
					    - apply E.
 | 
				
			||||||
    - intros a x. apply (U (L a) x).
 | 
					    - intros a x.
 | 
				
			||||||
 | 
					      apply ({|a|} ∪ x).
 | 
				
			||||||
    - intros. cbn.  
 | 
					    - intros. cbn.  
 | 
				
			||||||
      etransitivity. apply assoc.
 | 
					      etransitivity. apply assoc.
 | 
				
			||||||
      apply (ap (fun y => U y x)).
 | 
					      apply (ap (∪ x)).
 | 
				
			||||||
      apply idem.
 | 
					      apply idem.
 | 
				
			||||||
    - intros. cbn.
 | 
					    - intros. cbn.
 | 
				
			||||||
      etransitivity. apply assoc.
 | 
					      etransitivity. apply assoc.
 | 
				
			||||||
      etransitivity. refine (ap (fun y => U y x) _ ).
 | 
					      etransitivity. refine (ap (∪ x) _ ).
 | 
				
			||||||
      apply FSet.comm.
 | 
					      apply FSet.comm.
 | 
				
			||||||
      symmetry. 
 | 
					      symmetry. 
 | 
				
			||||||
      apply assoc.
 | 
					      apply assoc.
 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Definition FSet_to_FSetC: FSet A -> FSetC A :=
 | 
					  Definition FSet_to_FSetC: FSet A -> FSetC A.
 | 
				
			||||||
    FSet_rec A (FSetC A) (FSetC.trunc A) Nil singleton append append_assoc 
 | 
					  Proof.
 | 
				
			||||||
             append_comm append_nl append_nr singleton_idem.
 | 
					    hrecursion.
 | 
				
			||||||
 | 
					    - apply ∅.
 | 
				
			||||||
 | 
					    - intro a. apply {|a|}.
 | 
				
			||||||
 | 
					    - intros X Y. apply (X ∪ Y).
 | 
				
			||||||
 | 
					    - apply append_assoc.
 | 
				
			||||||
 | 
					    - apply append_comm.
 | 
				
			||||||
 | 
					    - apply append_nl.
 | 
				
			||||||
 | 
					    - apply append_nr.
 | 
				
			||||||
 | 
					    - apply singleton_idem.
 | 
				
			||||||
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Lemma append_union: forall (x y: FSetC A), 
 | 
					  Lemma append_union: forall (x y: FSetC A), 
 | 
				
			||||||
      FSetC_to_FSet (append x y) = U (FSetC_to_FSet x) (FSetC_to_FSet y).
 | 
					      FSetC_to_FSet (x ∪ y) = (FSetC_to_FSet x) ∪ (FSetC_to_FSet y).
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    intros x. 
 | 
					    intros x. 
 | 
				
			||||||
    hrecursion x; try (intros; apply path_forall; intro; apply set_path2).
 | 
					    hrecursion x; try (intros; apply path_forall; intro; apply set_path2).
 | 
				
			||||||
 
 | 
				
			|||||||
@@ -8,7 +8,7 @@ Proof.
 | 
				
			|||||||
  hrecursion.
 | 
					  hrecursion.
 | 
				
			||||||
  - exact ∅.
 | 
					  - exact ∅.
 | 
				
			||||||
  - intro a. exact {| f a |}.
 | 
					  - intro a. exact {| f a |}.
 | 
				
			||||||
  - exact U.
 | 
					  - intros X Y. apply (X ∪ Y).
 | 
				
			||||||
  - apply assoc.
 | 
					  - apply assoc.
 | 
				
			||||||
  - apply comm.
 | 
					  - apply comm.
 | 
				
			||||||
  - apply nl.
 | 
					  - apply nl.
 | 
				
			||||||
@@ -39,7 +39,7 @@ Proof.
 | 
				
			|||||||
hrecursion.
 | 
					hrecursion.
 | 
				
			||||||
- exact ∅.
 | 
					- exact ∅.
 | 
				
			||||||
- exact idmap.
 | 
					- exact idmap.
 | 
				
			||||||
- exact U.
 | 
					- intros X Y. apply (X ∪ Y).
 | 
				
			||||||
- apply assoc.
 | 
					- apply assoc.
 | 
				
			||||||
- apply comm.
 | 
					- apply comm.
 | 
				
			||||||
- apply nl.
 | 
					- apply nl.
 | 
				
			||||||
 
 | 
				
			|||||||
@@ -23,32 +23,6 @@ Section operations.
 | 
				
			|||||||
    - intros ; apply lor_idem.
 | 
					    - intros ; apply lor_idem.
 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Definition subset : FSet A -> FSet A -> hProp.
 | 
					 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    intros X Y.
 | 
					 | 
				
			||||||
    hrecursion X.
 | 
					 | 
				
			||||||
    - exists Unit.
 | 
					 | 
				
			||||||
      exact _.
 | 
					 | 
				
			||||||
    - intros a.
 | 
					 | 
				
			||||||
      apply (isIn a Y).
 | 
					 | 
				
			||||||
    - intros X1 X2.
 | 
					 | 
				
			||||||
      exists (prod X1 X2).
 | 
					 | 
				
			||||||
      exact _.
 | 
					 | 
				
			||||||
    - intros.
 | 
					 | 
				
			||||||
      apply path_trunctype ; apply equiv_prod_assoc.
 | 
					 | 
				
			||||||
    - intros.
 | 
					 | 
				
			||||||
      apply path_trunctype ; apply equiv_prod_symm.
 | 
					 | 
				
			||||||
    - intros.
 | 
					 | 
				
			||||||
      apply path_trunctype ; apply prod_unit_l.
 | 
					 | 
				
			||||||
    - intros.
 | 
					 | 
				
			||||||
      apply path_trunctype ; apply prod_unit_r.
 | 
					 | 
				
			||||||
    - intros a'.
 | 
					 | 
				
			||||||
      apply path_iff_hprop ; cbn.
 | 
					 | 
				
			||||||
      * intros [p1 p2]. apply p1.
 | 
					 | 
				
			||||||
      * intros p.
 | 
					 | 
				
			||||||
        split ; apply p.
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Definition comprehension : 
 | 
					  Definition comprehension : 
 | 
				
			||||||
    (A -> Bool) -> FSet A -> FSet A.
 | 
					    (A -> Bool) -> FSet A -> FSet A.
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
@@ -57,7 +31,7 @@ Section operations.
 | 
				
			|||||||
    - apply ∅.
 | 
					    - apply ∅.
 | 
				
			||||||
    - intro a.
 | 
					    - intro a.
 | 
				
			||||||
      refine (if (P a) then {|a|} else ∅).
 | 
					      refine (if (P a) then {|a|} else ∅).
 | 
				
			||||||
    - apply U.
 | 
					    - apply (∪).
 | 
				
			||||||
    - apply assoc.
 | 
					    - apply assoc.
 | 
				
			||||||
    - apply comm.
 | 
					    - apply comm.
 | 
				
			||||||
    - apply nl.
 | 
					    - apply nl.
 | 
				
			||||||
@@ -82,7 +56,7 @@ Section operations.
 | 
				
			|||||||
    - apply ∅.
 | 
					    - apply ∅.
 | 
				
			||||||
    - intro b.
 | 
					    - intro b.
 | 
				
			||||||
      apply {|(a, b)|}.
 | 
					      apply {|(a, b)|}.
 | 
				
			||||||
    - apply U.
 | 
					    - apply (∪).
 | 
				
			||||||
    - intros X Y Z ; apply assoc.
 | 
					    - intros X Y Z ; apply assoc.
 | 
				
			||||||
    - intros X Y ; apply comm.
 | 
					    - intros X Y ; apply comm.
 | 
				
			||||||
    - intros ; apply nl.
 | 
					    - intros ; apply nl.
 | 
				
			||||||
@@ -97,7 +71,7 @@ Section operations.
 | 
				
			|||||||
    - apply ∅.
 | 
					    - apply ∅.
 | 
				
			||||||
    - intro a.
 | 
					    - intro a.
 | 
				
			||||||
      apply (single_product a Y).
 | 
					      apply (single_product a Y).
 | 
				
			||||||
    - apply U.
 | 
					    - apply (∪).
 | 
				
			||||||
    - intros ; apply assoc.
 | 
					    - intros ; apply assoc.
 | 
				
			||||||
    - intros ; apply comm.
 | 
					    - intros ; apply comm.
 | 
				
			||||||
    - intros ; apply nl.
 | 
					    - intros ; apply nl.
 | 
				
			||||||
@@ -107,5 +81,48 @@ Section operations.
 | 
				
			|||||||
  
 | 
					  
 | 
				
			||||||
End operations.
 | 
					End operations.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
Infix "∈" := isIn (at level 9, right associativity).
 | 
					Section instances_operations.
 | 
				
			||||||
 | 
					  Global Instance fset_comprehension : hasComprehension FSet.
 | 
				
			||||||
 | 
					  Proof.
 | 
				
			||||||
 | 
					    intros A ϕ X.
 | 
				
			||||||
 | 
					    apply (comprehension ϕ X).
 | 
				
			||||||
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					  Context `{Univalence}.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					  Global Instance fset_member : hasMembership FSet.
 | 
				
			||||||
 | 
					  Proof.
 | 
				
			||||||
 | 
					    intros A a X.
 | 
				
			||||||
 | 
					    apply (isIn a X).
 | 
				
			||||||
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					  Context {A : Type}.
 | 
				
			||||||
 | 
					  
 | 
				
			||||||
 | 
					  Definition subset : FSet A -> FSet A -> hProp.
 | 
				
			||||||
 | 
					  Proof.
 | 
				
			||||||
 | 
					    intros X Y.
 | 
				
			||||||
 | 
					    hrecursion X.
 | 
				
			||||||
 | 
					    - exists Unit.
 | 
				
			||||||
 | 
					      exact _.
 | 
				
			||||||
 | 
					    - intros a.
 | 
				
			||||||
 | 
					      apply (a ∈ Y).
 | 
				
			||||||
 | 
					    - intros X1 X2.
 | 
				
			||||||
 | 
					      exists (prod X1 X2).
 | 
				
			||||||
 | 
					      exact _.
 | 
				
			||||||
 | 
					    - intros.
 | 
				
			||||||
 | 
					      apply path_trunctype ; apply equiv_prod_assoc.
 | 
				
			||||||
 | 
					    - intros.
 | 
				
			||||||
 | 
					      apply path_trunctype ; apply equiv_prod_symm.
 | 
				
			||||||
 | 
					    - intros.
 | 
				
			||||||
 | 
					      apply path_trunctype ; apply prod_unit_l.
 | 
				
			||||||
 | 
					    - intros.
 | 
				
			||||||
 | 
					      apply path_trunctype ; apply prod_unit_r.
 | 
				
			||||||
 | 
					    - intros a'.
 | 
				
			||||||
 | 
					      apply path_iff_hprop ; cbn.
 | 
				
			||||||
 | 
					      * intros [p1 p2]. apply p1.
 | 
				
			||||||
 | 
					      * intros p.
 | 
				
			||||||
 | 
					        split ; apply p.
 | 
				
			||||||
 | 
					  Defined.
 | 
				
			||||||
 | 
					End instances_operations.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
Infix  "⊆" := subset (at level 10, right associativity).
 | 
					Infix  "⊆" := subset (at level 10, right associativity).
 | 
				
			||||||
@@ -4,9 +4,9 @@ Require Import representations.cons_repr.
 | 
				
			|||||||
 | 
					
 | 
				
			||||||
Section operations.
 | 
					Section operations.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Context {A : Type}.
 | 
					  Global Instance fsetc_union : hasUnion FSetC.
 | 
				
			||||||
 | 
					  Proof.
 | 
				
			||||||
  Definition append  (x y: FSetC A) : FSetC A.
 | 
					    intros A x y.
 | 
				
			||||||
    hinduction x.
 | 
					    hinduction x.
 | 
				
			||||||
    - apply y.
 | 
					    - apply y.
 | 
				
			||||||
    - apply Cns.
 | 
					    - apply Cns.
 | 
				
			||||||
@@ -14,6 +14,6 @@ Section operations.
 | 
				
			|||||||
    - apply comm.
 | 
					    - apply comm.
 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Definition singleton (a:A) : FSetC A := a;;∅.
 | 
					  Global Instance fsetc_singleton : hasSingleton FSetC := fun A a => a;;∅.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
End operations.
 | 
					End operations.
 | 
				
			||||||
@@ -33,7 +33,7 @@ Section decidable_A.
 | 
				
			|||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    hinduction ; try (intros ; apply path_ishprop).
 | 
					    hinduction ; try (intros ; apply path_ishprop).
 | 
				
			||||||
    - intro ; apply _.
 | 
					    - intro ; apply _.
 | 
				
			||||||
    - intros ; apply _. 
 | 
					    - intros. apply _. 
 | 
				
			||||||
    - intros ; apply _.
 | 
					    - intros ; apply _.
 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 
 | 
				
			|||||||
@@ -9,17 +9,23 @@ Section characterize_isIn.
 | 
				
			|||||||
  Context `{Univalence}.
 | 
					  Context `{Univalence}.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  (** isIn properties *)
 | 
					  (** isIn properties *)
 | 
				
			||||||
  Definition empty_isIn (a: A) : a ∈ E -> Empty := idmap.
 | 
					  Definition empty_isIn (a: A) : a ∈ ∅ -> Empty := idmap.
 | 
				
			||||||
  
 | 
					  
 | 
				
			||||||
  Definition singleton_isIn (a b: A) : a ∈ {|b|} -> Trunc (-1) (a = b) := idmap.
 | 
					  Definition singleton_isIn (a b: A) : a ∈ {|b|} -> Trunc (-1) (a = b) := idmap.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Definition union_isIn (X Y : FSet A) (a : A)
 | 
					  Definition union_isIn (X Y : FSet A) (a : A)
 | 
				
			||||||
    : a ∈ X ∪ Y = a ∈ X ∨ a ∈ Y := idpath.
 | 
					    : a ∈ X ∪ Y = a ∈ X ∨ a ∈ Y := idpath.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Lemma comprehension_isIn (ϕ : A -> Bool) (a : A) : forall X : FSet A,
 | 
					  Lemma comprehension_union (ϕ : A -> Bool) : forall X Y : FSet A,
 | 
				
			||||||
      a ∈ (comprehension ϕ X) = if ϕ a then a ∈ X else False_hp.
 | 
					      {|X ∪ Y & ϕ|} = {|X & ϕ|} ∪ {|Y & ϕ|}.
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    hinduction ; try (intros ; apply set_path2) ; cbn.
 | 
					    reflexivity.
 | 
				
			||||||
 | 
					  Defined.  
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					  Lemma comprehension_isIn (ϕ : A -> Bool) (a : A) : forall X : FSet A,
 | 
				
			||||||
 | 
					      a ∈ {|X & ϕ|} = if ϕ a then a ∈ X else False_hp.
 | 
				
			||||||
 | 
					  Proof.
 | 
				
			||||||
 | 
					    hinduction ; try (intros ; apply set_path2).
 | 
				
			||||||
    - destruct (ϕ a) ; reflexivity.
 | 
					    - destruct (ϕ a) ; reflexivity.
 | 
				
			||||||
    - intros b.
 | 
					    - intros b.
 | 
				
			||||||
      assert (forall c d, ϕ a = c -> ϕ b = d ->
 | 
					      assert (forall c d, ϕ a = c -> ϕ b = d ->
 | 
				
			||||||
@@ -36,6 +42,8 @@ Section characterize_isIn.
 | 
				
			|||||||
      }
 | 
					      }
 | 
				
			||||||
      apply (X (ϕ a) (ϕ b) idpath idpath).
 | 
					      apply (X (ϕ a) (ϕ b) idpath idpath).
 | 
				
			||||||
    - intros X Y H1 H2.
 | 
					    - intros X Y H1 H2.
 | 
				
			||||||
 | 
					      rewrite comprehension_union.
 | 
				
			||||||
 | 
					      rewrite union_isIn.
 | 
				
			||||||
      rewrite H1, H2.
 | 
					      rewrite H1, H2.
 | 
				
			||||||
      destruct (ϕ a).
 | 
					      destruct (ϕ a).
 | 
				
			||||||
      * reflexivity.
 | 
					      * reflexivity.
 | 
				
			||||||
@@ -44,11 +52,14 @@ Section characterize_isIn.
 | 
				
			|||||||
           destruct Z ; assumption.
 | 
					           destruct Z ; assumption.
 | 
				
			||||||
        ** intros ; apply tr ; right ; assumption.
 | 
					        ** intros ; apply tr ; right ; assumption.
 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					End characterize_isIn.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Context {B : Type}.
 | 
					Section product_isIn.
 | 
				
			||||||
 | 
					  Context {A B : Type}.
 | 
				
			||||||
 | 
					  Context `{Univalence}.
 | 
				
			||||||
  
 | 
					  
 | 
				
			||||||
  Lemma isIn_singleproduct (a : A) (b : B) (c : A) : forall (Y : FSet B),
 | 
					  Lemma isIn_singleproduct (a : A) (b : B) (c : A) : forall (Y : FSet B),
 | 
				
			||||||
      isIn (a, b) (single_product c Y) = land (BuildhProp (Trunc (-1) (a = c))) (isIn b Y).
 | 
					      (a, b) ∈ (single_product c Y) = land (BuildhProp (Trunc (-1) (a = c))) (b ∈ Y).
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    hinduction ; try (intros ; apply path_ishprop).
 | 
					    hinduction ; try (intros ; apply path_ishprop).
 | 
				
			||||||
    - apply path_hprop ; symmetry ; apply prod_empty_r.
 | 
					    - apply path_hprop ; symmetry ; apply prod_empty_r.
 | 
				
			||||||
@@ -62,16 +73,16 @@ Section characterize_isIn.
 | 
				
			|||||||
        rewrite Z1, Z2.
 | 
					        rewrite Z1, Z2.
 | 
				
			||||||
        apply (tr idpath).
 | 
					        apply (tr idpath).
 | 
				
			||||||
    - intros X1 X2 HX1 HX2.
 | 
					    - intros X1 X2 HX1 HX2.
 | 
				
			||||||
      apply path_iff_hprop.
 | 
					      apply path_iff_hprop ; rewrite ?union_isIn.
 | 
				
			||||||
      *  intros X.
 | 
					      * intros X.
 | 
				
			||||||
         strip_truncations.
 | 
					        strip_truncations.
 | 
				
			||||||
         destruct X as [H1 | H1] ; rewrite ?HX1, ?HX2 in H1 ; destruct H1 as [H1 H2].
 | 
					        destruct X as [H1 | H1] ; rewrite ?HX1, ?HX2 in H1 ; destruct H1 as [H1 H2].
 | 
				
			||||||
         ** split.
 | 
					        ** split.
 | 
				
			||||||
            *** apply H1.
 | 
					           *** apply H1.
 | 
				
			||||||
            *** apply (tr(inl H2)).
 | 
					           *** apply (tr(inl H2)).
 | 
				
			||||||
         ** split.
 | 
					        ** split.
 | 
				
			||||||
            *** apply H1.
 | 
					           *** apply H1.
 | 
				
			||||||
            *** apply (tr(inr H2)).
 | 
					           *** apply (tr(inr H2)).
 | 
				
			||||||
      * intros [H1 H2].
 | 
					      * intros [H1 H2].
 | 
				
			||||||
        strip_truncations.
 | 
					        strip_truncations.
 | 
				
			||||||
        apply tr.
 | 
					        apply tr.
 | 
				
			||||||
@@ -84,15 +95,16 @@ Section characterize_isIn.
 | 
				
			|||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
  
 | 
					  
 | 
				
			||||||
  Definition isIn_product (a : A) (b : B) (X : FSet A) (Y : FSet B) :
 | 
					  Definition isIn_product (a : A) (b : B) (X : FSet A) (Y : FSet B) :
 | 
				
			||||||
    isIn (a,b) (product X Y) = land (isIn a X) (isIn b Y).
 | 
					    (a,b) ∈ (product X Y) = land (a ∈ X) (b ∈ Y).
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    hinduction X ; try (intros ; apply path_ishprop).
 | 
					    hinduction X ; try (intros ; apply path_ishprop).
 | 
				
			||||||
    - apply path_hprop ; symmetry ; apply prod_empty_l.
 | 
					    - apply path_hprop ; symmetry ; apply prod_empty_l.
 | 
				
			||||||
    - intros.
 | 
					    - intros.
 | 
				
			||||||
      apply isIn_singleproduct.
 | 
					      apply isIn_singleproduct.
 | 
				
			||||||
    - intros X1 X2 HX1 HX2.
 | 
					    - intros X1 X2 HX1 HX2.
 | 
				
			||||||
 | 
					      rewrite (union_isIn (product X1 Y)).
 | 
				
			||||||
      rewrite HX1, HX2.
 | 
					      rewrite HX1, HX2.
 | 
				
			||||||
      apply path_iff_hprop.
 | 
					      apply path_iff_hprop ; rewrite ?union_isIn.
 | 
				
			||||||
      * intros X.
 | 
					      * intros X.
 | 
				
			||||||
        strip_truncations.
 | 
					        strip_truncations.
 | 
				
			||||||
        destruct X as [[H3 H4] | [H3 H4]] ; split ; try (apply H4).
 | 
					        destruct X as [[H3 H4] | [H3 H4]] ; split ; try (apply H4).
 | 
				
			||||||
@@ -104,7 +116,7 @@ Section characterize_isIn.
 | 
				
			|||||||
        ** left ; split ; assumption.
 | 
					        ** left ; split ; assumption.
 | 
				
			||||||
        ** right ; split ; assumption.
 | 
					        ** right ; split ; assumption.
 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
End characterize_isIn.
 | 
					End product_isIn.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
Ltac simplify_isIn :=
 | 
					Ltac simplify_isIn :=
 | 
				
			||||||
  repeat (rewrite union_isIn
 | 
					  repeat (rewrite union_isIn
 | 
				
			||||||
@@ -120,8 +132,17 @@ Section properties.
 | 
				
			|||||||
  Context {A : Type}.
 | 
					  Context {A : Type}.
 | 
				
			||||||
  Context `{Univalence}.
 | 
					  Context `{Univalence}.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Instance: bottom(FSet A) := ∅.
 | 
					  Instance: bottom(FSet A).
 | 
				
			||||||
  Instance: maximum(FSet A) := U.
 | 
					  Proof.
 | 
				
			||||||
 | 
					    unfold bottom.
 | 
				
			||||||
 | 
					    apply ∅.
 | 
				
			||||||
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					  Instance: maximum(FSet A).
 | 
				
			||||||
 | 
					  Proof.
 | 
				
			||||||
 | 
					    intros x y.
 | 
				
			||||||
 | 
					    apply (x ∪ y).
 | 
				
			||||||
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Global Instance joinsemilattice_fset : JoinSemiLattice (FSet A).
 | 
					  Global Instance joinsemilattice_fset : JoinSemiLattice (FSet A).
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
@@ -129,27 +150,33 @@ Section properties.
 | 
				
			|||||||
  Defined.  
 | 
					  Defined.  
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  (** comprehension properties *)
 | 
					  (** comprehension properties *)
 | 
				
			||||||
  Lemma comprehension_false Y : comprehension (fun (_ : A) => false) Y = ∅.
 | 
					  Lemma comprehension_false X : {|X & (fun (_ : A) => false)|} = ∅.
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    toHProp.
 | 
					    toHProp.
 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Lemma comprehension_subset : forall ϕ (X : FSet A),
 | 
					  Lemma comprehension_subset : forall ϕ (X : FSet A),
 | 
				
			||||||
      (comprehension ϕ X) ∪ X = X.
 | 
					      {|X & ϕ|} ∪ X = X.
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    toHProp.
 | 
					    toHProp.
 | 
				
			||||||
    destruct (ϕ a) ; eauto with lattice_hints typeclass_instances.
 | 
					    destruct (ϕ a) ; eauto with lattice_hints typeclass_instances.
 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Lemma comprehension_or : forall ϕ ψ (x: FSet A),
 | 
					  Lemma comprehension_or : forall ϕ ψ (X : FSet A),
 | 
				
			||||||
      comprehension (fun a => orb (ϕ a) (ψ a)) x
 | 
					      {|X & (fun a => orb (ϕ a) (ψ a))|}
 | 
				
			||||||
      = (comprehension ϕ x) ∪ (comprehension ψ x).
 | 
					      = {|X & ϕ|} ∪ {|X & ψ|}.
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    toHProp.
 | 
					    toHProp.
 | 
				
			||||||
    symmetry ; destruct (ϕ a) ; destruct (ψ a)
 | 
					    symmetry ; destruct (ϕ a) ; destruct (ψ a)
 | 
				
			||||||
    ; eauto with lattice_hints typeclass_instances.
 | 
					    ; eauto with lattice_hints typeclass_instances.
 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					  Lemma comprehension_all : forall (X : FSet A),
 | 
				
			||||||
 | 
					      comprehension (fun a => true) X = X.
 | 
				
			||||||
 | 
					  Proof.
 | 
				
			||||||
 | 
					    toHProp.
 | 
				
			||||||
 | 
					  Defined.
 | 
				
			||||||
 | 
					  
 | 
				
			||||||
  Lemma merely_choice : forall X : FSet A, hor (X = ∅) (hexists (fun a => a ∈ X)).
 | 
					  Lemma merely_choice : forall X : FSet A, hor (X = ∅) (hexists (fun a => a ∈ X)).
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    hinduction; try (intros; apply equiv_hprop_allpath ; apply _).
 | 
					    hinduction; try (intros; apply equiv_hprop_allpath ; apply _).
 | 
				
			||||||
@@ -161,7 +188,7 @@ Section properties.
 | 
				
			|||||||
      destruct TX as [XE | HX] ; destruct TY as [YE | HY] ; try(strip_truncations ; apply tr).
 | 
					      destruct TX as [XE | HX] ; destruct TY as [YE | HY] ; try(strip_truncations ; apply tr).
 | 
				
			||||||
      * refine (tr (inl _)).
 | 
					      * refine (tr (inl _)).
 | 
				
			||||||
        rewrite XE, YE.
 | 
					        rewrite XE, YE.
 | 
				
			||||||
        apply (union_idem E).
 | 
					        apply (union_idem ∅).
 | 
				
			||||||
      * destruct HY as [a Ya].
 | 
					      * destruct HY as [a Ya].
 | 
				
			||||||
        refine (inr (tr _)).
 | 
					        refine (inr (tr _)).
 | 
				
			||||||
        exists a.
 | 
					        exists a.
 | 
				
			||||||
 
 | 
				
			|||||||
@@ -6,38 +6,30 @@ From fsets Require Import operations_cons_repr.
 | 
				
			|||||||
Section properties.
 | 
					Section properties.
 | 
				
			||||||
  Context {A : Type}.
 | 
					  Context {A : Type}.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Lemma append_nl: 
 | 
					  Definition append_nl : forall (x: FSetC A), ∅ ∪ x = x
 | 
				
			||||||
    forall (x: FSetC A), append ∅ x = x. 
 | 
					    := fun _ => idpath. 
 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    intro. reflexivity.
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Lemma append_nr: 
 | 
					  Lemma append_nr : forall (x: FSetC A), x ∪ ∅ = x.
 | 
				
			||||||
    forall (x: FSetC A), append x ∅ = x.
 | 
					 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    hinduction; try (intros; apply set_path2).
 | 
					    hinduction; try (intros; apply set_path2).
 | 
				
			||||||
    -  reflexivity.
 | 
					    -  reflexivity.
 | 
				
			||||||
    -  intros. apply (ap (fun y => a ;; y) X).
 | 
					    -  intros. apply (ap (fun y => a;;y) X).
 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
  
 | 
					  
 | 
				
			||||||
  Lemma append_assoc {H: Funext}:  
 | 
					  Lemma append_assoc {H: Funext}:  
 | 
				
			||||||
    forall (x y z: FSetC A), append x (append y z) = append (append x y) z.
 | 
					    forall (x y z: FSetC A), x ∪ (y ∪ z) = (x ∪ y) ∪ z.
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    intro x; hinduction x; try (intros; apply set_path2).
 | 
					    hinduction
 | 
				
			||||||
 | 
					    ; try (intros ; apply path_forall ; intro
 | 
				
			||||||
 | 
					           ; apply path_forall ; intro ;  apply set_path2).
 | 
				
			||||||
    - reflexivity.
 | 
					    - reflexivity.
 | 
				
			||||||
    - intros a x HR y z. 
 | 
					    - intros a x HR y z. 
 | 
				
			||||||
      specialize (HR y z).
 | 
					      specialize (HR y z).
 | 
				
			||||||
      apply (ap (fun y => a ;; y) HR). 
 | 
					      apply (ap (fun y => a;;y) HR). 
 | 
				
			||||||
    - intros. apply path_forall. 
 | 
					 | 
				
			||||||
      intro. apply path_forall. 
 | 
					 | 
				
			||||||
      intro. apply set_path2.
 | 
					 | 
				
			||||||
    - intros. apply path_forall.
 | 
					 | 
				
			||||||
      intro. apply path_forall. 
 | 
					 | 
				
			||||||
      intro. apply set_path2.
 | 
					 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
  
 | 
					  
 | 
				
			||||||
  Lemma append_singleton: forall (a: A) (x: FSetC A), 
 | 
					  Lemma append_singleton: forall (a: A) (x: FSetC A), 
 | 
				
			||||||
      a ;; x = append x (a ;; ∅).
 | 
					      a ;; x = x ∪ (a ;; ∅).
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    intro a. hinduction; try (intros; apply set_path2).
 | 
					    intro a. hinduction; try (intros; apply set_path2).
 | 
				
			||||||
    - reflexivity.
 | 
					    - reflexivity.
 | 
				
			||||||
@@ -47,29 +39,26 @@ Section properties.
 | 
				
			|||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Lemma append_comm {H: Funext}: 
 | 
					  Lemma append_comm {H: Funext}: 
 | 
				
			||||||
    forall (x1 x2: FSetC A), append x1 x2 = append x2 x1.
 | 
					    forall (x1 x2: FSetC A), x1 ∪ x2 = x2 ∪ x1.
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    intro x1. 
 | 
					    hinduction ;  try (intros ; apply path_forall ; intro ; apply set_path2).
 | 
				
			||||||
    hinduction x1;  try (intros; apply set_path2).
 | 
					 | 
				
			||||||
    - intros. symmetry. apply append_nr. 
 | 
					    - intros. symmetry. apply append_nr. 
 | 
				
			||||||
    - intros a x1 HR x2.
 | 
					    - intros a x1 HR x2.
 | 
				
			||||||
      etransitivity.
 | 
					      etransitivity.
 | 
				
			||||||
      apply (ap (fun y => a;;y) (HR x2)).  
 | 
					      apply (ap (fun y => a;;y) (HR x2)).  
 | 
				
			||||||
      transitivity  (append (append x2 x1) (a;;∅)).
 | 
					      transitivity  ((x2 ∪ x1) ∪ (a;;∅)).
 | 
				
			||||||
      + apply append_singleton. 
 | 
					      + apply append_singleton. 
 | 
				
			||||||
      + etransitivity.
 | 
					      + etransitivity.
 | 
				
			||||||
    	* symmetry. apply append_assoc.
 | 
					    	* symmetry. apply append_assoc.
 | 
				
			||||||
    	* simple refine (ap (fun y => append x2 y) (append_singleton _ _)^).
 | 
					    	* simple refine (ap (x2 ∪) (append_singleton _ _)^).
 | 
				
			||||||
    - intros. apply path_forall.
 | 
					 | 
				
			||||||
      intro. apply set_path2.
 | 
					 | 
				
			||||||
    - intros. apply path_forall.
 | 
					 | 
				
			||||||
      intro. apply set_path2.  
 | 
					 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Lemma singleton_idem: forall (a: A), 
 | 
					  Lemma singleton_idem: forall (a: A), 
 | 
				
			||||||
      append (singleton a) (singleton a) = singleton a.
 | 
					      {|a|} ∪ {|a|} = {|a|}.
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    unfold singleton. intro. cbn. apply dupl.
 | 
					    unfold singleton.
 | 
				
			||||||
 | 
					    intro.
 | 
				
			||||||
 | 
					    apply dupl.
 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
End properties.
 | 
					End properties.
 | 
				
			||||||
 
 | 
				
			|||||||
@@ -71,10 +71,10 @@ Section operations_isIn.
 | 
				
			|||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Lemma comprehension_isIn_b (Y : FSet A) (ϕ : A -> Bool) (a : A) :
 | 
					  Lemma comprehension_isIn_b (Y : FSet A) (ϕ : A -> Bool) (a : A) :
 | 
				
			||||||
    isIn_b a (comprehension ϕ Y) = andb (isIn_b a Y) (ϕ a).
 | 
					    isIn_b a {|Y & ϕ|} = andb (isIn_b a Y) (ϕ a).
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    unfold isIn_b, dec ; simpl.
 | 
					    unfold isIn_b, dec ; simpl.
 | 
				
			||||||
    destruct (isIn_decidable a (comprehension ϕ Y)) as [t | t]
 | 
					    destruct (isIn_decidable a {|Y & ϕ|}) as [t | t]
 | 
				
			||||||
    ; destruct (isIn_decidable a Y) as [n | n] ; rewrite comprehension_isIn in t
 | 
					    ; destruct (isIn_decidable a Y) as [n | n] ; rewrite comprehension_isIn in t
 | 
				
			||||||
    ; destruct (ϕ a) ; try reflexivity ; try contradiction.
 | 
					    ; destruct (ϕ a) ; try reflexivity ; try contradiction.
 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
@@ -105,9 +105,23 @@ Section SetLattice.
 | 
				
			|||||||
  Context {A_deceq : DecidablePaths A}.
 | 
					  Context {A_deceq : DecidablePaths A}.
 | 
				
			||||||
  Context `{Univalence}.
 | 
					  Context `{Univalence}.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Instance fset_max : maximum (FSet A) := U.
 | 
					  Instance fset_max : maximum (FSet A).
 | 
				
			||||||
  Instance fset_min : minimum (FSet A) := intersection.
 | 
					  Proof.
 | 
				
			||||||
  Instance fset_bot : bottom (FSet A) := ∅.
 | 
					    intros x y.
 | 
				
			||||||
 | 
					    apply (x ∪ y).
 | 
				
			||||||
 | 
					  Defined.
 | 
				
			||||||
 | 
					      
 | 
				
			||||||
 | 
					  Instance fset_min : minimum (FSet A).
 | 
				
			||||||
 | 
					  Proof.
 | 
				
			||||||
 | 
					    intros x y.
 | 
				
			||||||
 | 
					    apply (intersection x y).
 | 
				
			||||||
 | 
					  Defined.
 | 
				
			||||||
 | 
					  
 | 
				
			||||||
 | 
					  Instance fset_bot : bottom (FSet A).
 | 
				
			||||||
 | 
					  Proof.
 | 
				
			||||||
 | 
					    unfold bottom.
 | 
				
			||||||
 | 
					    apply ∅.
 | 
				
			||||||
 | 
					  Defined.
 | 
				
			||||||
    
 | 
					    
 | 
				
			||||||
  Instance lattice_fset : Lattice (FSet A).
 | 
					  Instance lattice_fset : Lattice (FSet A).
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
@@ -116,40 +130,6 @@ Section SetLattice.
 | 
				
			|||||||
  
 | 
					  
 | 
				
			||||||
End SetLattice.
 | 
					End SetLattice.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
(* Comprehension properties *)
 | 
					 | 
				
			||||||
Section comprehension_properties.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Context {A : Type}.
 | 
					 | 
				
			||||||
  Context {A_deceq : DecidablePaths A}.
 | 
					 | 
				
			||||||
  Context `{Univalence}.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Lemma comprehension_or : forall ϕ ψ (x: FSet A),
 | 
					 | 
				
			||||||
      comprehension (fun a => orb (ϕ a) (ψ a)) x
 | 
					 | 
				
			||||||
      = U (comprehension ϕ x) (comprehension ψ x).
 | 
					 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    toBool.
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
  
 | 
					 | 
				
			||||||
  (** comprehension properties *)
 | 
					 | 
				
			||||||
  Lemma comprehension_false Y : comprehension (fun (_ : A) => false) Y = ∅.
 | 
					 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    toBool.
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Lemma comprehension_all : forall (X : FSet A),
 | 
					 | 
				
			||||||
      comprehension (fun a => isIn_b a X) X = X.
 | 
					 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    toBool.
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
  
 | 
					 | 
				
			||||||
  Lemma comprehension_subset : forall ϕ (X : FSet A),
 | 
					 | 
				
			||||||
      (comprehension ϕ X) ∪ X = X.
 | 
					 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    toBool.
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
  
 | 
					 | 
				
			||||||
End comprehension_properties.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
(* With extensionality we can prove decidable equality *)
 | 
					(* With extensionality we can prove decidable equality *)
 | 
				
			||||||
Section dec_eq.
 | 
					Section dec_eq.
 | 
				
			||||||
  Context (A : Type) `{DecidablePaths A} `{Univalence}.
 | 
					  Context (A : Type) `{DecidablePaths A} `{Univalence}.
 | 
				
			||||||
 
 | 
				
			|||||||
@@ -1,31 +1,6 @@
 | 
				
			|||||||
Require Import HoTT.
 | 
					Require Import HoTT.
 | 
				
			||||||
Require Import FSets.
 | 
					Require Import FSets.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
Section structure.
 | 
					 | 
				
			||||||
  Variable (T : Type -> Type).
 | 
					 | 
				
			||||||
  
 | 
					 | 
				
			||||||
  Class hasMembership : Type :=
 | 
					 | 
				
			||||||
    member : forall A : Type, A -> T A -> hProp.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Class hasEmpty : Type :=
 | 
					 | 
				
			||||||
    empty : forall A, T A.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Class hasSingleton : Type :=
 | 
					 | 
				
			||||||
    singleton : forall A, A -> T A.
 | 
					 | 
				
			||||||
  
 | 
					 | 
				
			||||||
  Class hasUnion : Type :=
 | 
					 | 
				
			||||||
    union : forall A, T A -> T A -> T A.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Class hasComprehension : Type :=
 | 
					 | 
				
			||||||
    filter : forall A, (A -> Bool) -> T A -> T A.
 | 
					 | 
				
			||||||
End structure.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
Arguments member {_} {_} {_} _ _.
 | 
					 | 
				
			||||||
Arguments empty {_} {_} {_}.
 | 
					 | 
				
			||||||
Arguments singleton {_} {_} {_} _.
 | 
					 | 
				
			||||||
Arguments union {_} {_} {_} _ _.
 | 
					 | 
				
			||||||
Arguments filter {_} {_} {_} _ _.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
Section interface.
 | 
					Section interface.
 | 
				
			||||||
  Context `{Univalence}.
 | 
					  Context `{Univalence}.
 | 
				
			||||||
  Variable (T : Type -> Type)
 | 
					  Variable (T : Type -> Type)
 | 
				
			||||||
 
 | 
				
			|||||||
@@ -1,16 +1,15 @@
 | 
				
			|||||||
(* Typeclass for lattices *)
 | 
					(* Typeclass for lattices *)
 | 
				
			||||||
Require Import HoTT.
 | 
					Require Import HoTT.
 | 
				
			||||||
 | 
					Require Import notation.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
Section binary_operation.
 | 
					Section binary_operation.
 | 
				
			||||||
  Variable A : Type.
 | 
					  Variable A : Type.
 | 
				
			||||||
  
 | 
					  
 | 
				
			||||||
  Definition operation := A -> A -> A.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Class maximum :=
 | 
					  Class maximum :=
 | 
				
			||||||
    max_L : operation.
 | 
					    max_L : operation A.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Class minimum :=
 | 
					  Class minimum :=
 | 
				
			||||||
    min_L : operation.
 | 
					    min_L : operation A.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Class bottom :=
 | 
					  Class bottom :=
 | 
				
			||||||
    empty : A.
 | 
					    empty : A.
 | 
				
			||||||
@@ -20,41 +19,6 @@ Arguments max_L {_} {_} _.
 | 
				
			|||||||
Arguments min_L {_} {_} _.
 | 
					Arguments min_L {_} {_} _.
 | 
				
			||||||
Arguments empty {_}.
 | 
					Arguments empty {_}.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
Section Defs.
 | 
					 | 
				
			||||||
  Variable A : Type.
 | 
					 | 
				
			||||||
  Variable f : A -> A -> A.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Class Commutative :=
 | 
					 | 
				
			||||||
    commutative : forall x y, f x y = f y x.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Class Associative :=
 | 
					 | 
				
			||||||
    associativity : forall x y z, f (f x y) z = f x (f y z).
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Class Idempotent :=
 | 
					 | 
				
			||||||
    idempotency : forall x, f x x = x.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Variable g : operation A.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Class Absorption :=
 | 
					 | 
				
			||||||
    absorb : forall x y, f x (g x y) = x.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Variable n : A.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Class NeutralL :=
 | 
					 | 
				
			||||||
    neutralityL : forall x, f n x = x.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Class NeutralR :=
 | 
					 | 
				
			||||||
    neutralityR : forall x, f x n = x.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
End Defs.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
Arguments Commutative {_} _.
 | 
					 | 
				
			||||||
Arguments Associative {_} _.
 | 
					 | 
				
			||||||
Arguments Idempotent {_} _.
 | 
					 | 
				
			||||||
Arguments NeutralL {_} _ _.
 | 
					 | 
				
			||||||
Arguments NeutralR {_} _ _.
 | 
					 | 
				
			||||||
Arguments Absorption {_} _ _.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
Section JoinSemiLattice.
 | 
					Section JoinSemiLattice.
 | 
				
			||||||
  Variable A : Type.
 | 
					  Variable A : Type.
 | 
				
			||||||
  Context {max_L : maximum A} {empty_L : bottom A}.
 | 
					  Context {max_L : maximum A} {empty_L : bottom A}.
 | 
				
			||||||
 
 | 
				
			|||||||
@@ -1,50 +1,45 @@
 | 
				
			|||||||
(* This is a /bad/ definition of FSets, without the 0-truncation.
 | 
					(* This is a /bad/ definition of FSets, without the 0-truncation.
 | 
				
			||||||
   Here we show that the resulting type is not an h-set. *)
 | 
					   Here we show that the resulting type is not an h-set. *)
 | 
				
			||||||
 | 
					
 | 
				
			||||||
Require Import HoTT.
 | 
					Require Import HoTT HitTactics.
 | 
				
			||||||
Require Import HitTactics.
 | 
					Require Import notation.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
Module Export FSet.
 | 
					Module Export FSet.
 | 
				
			||||||
  
 | 
					  
 | 
				
			||||||
  Section FSet.
 | 
					  Section FSet.
 | 
				
			||||||
 | 
					    Private Inductive FSet (A : Type): Type :=
 | 
				
			||||||
 | 
					    | E : FSet A
 | 
				
			||||||
 | 
					    | L : A -> FSet A
 | 
				
			||||||
 | 
					    | U : FSet A -> FSet A -> FSet A.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					    Global Instance fset_empty : hasEmpty FSet := E.
 | 
				
			||||||
 | 
					    Global Instance fset_singleton : hasSingleton FSet := L.
 | 
				
			||||||
 | 
					    Global Instance fset_union : hasUnion FSet := U.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
    Variable A : Type.
 | 
					    Variable A : Type.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
    Private Inductive FSet : Type :=
 | 
					    Axiom assoc : forall (x y z : FSet A),
 | 
				
			||||||
    | E : FSet
 | 
					 | 
				
			||||||
    | L : A -> FSet
 | 
					 | 
				
			||||||
    | U : FSet -> FSet -> FSet.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
    Notation "{| x |}" :=  (L x).
 | 
					 | 
				
			||||||
    Infix "∪" := U (at level 8, right associativity).
 | 
					 | 
				
			||||||
    Notation "∅" := E.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
    Axiom assoc : forall (x y z : FSet ),
 | 
					 | 
				
			||||||
        x ∪ (y ∪ z) = (x ∪ y) ∪ z.
 | 
					        x ∪ (y ∪ z) = (x ∪ y) ∪ z.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
    Axiom comm : forall (x y : FSet),
 | 
					    Axiom comm : forall (x y : FSet A),
 | 
				
			||||||
        x ∪ y = y ∪ x.
 | 
					        x ∪ y = y ∪ x.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
    Axiom nl : forall (x : FSet),
 | 
					    Axiom nl : forall (x : FSet A),
 | 
				
			||||||
        ∅ ∪ x = x.
 | 
					        ∅ ∪ x = x.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
    Axiom nr : forall (x : FSet),
 | 
					    Axiom nr : forall (x : FSet A),
 | 
				
			||||||
        x ∪ ∅ = x.
 | 
					        x ∪ ∅ = x.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
    Axiom idem : forall (x : A),
 | 
					    Axiom idem : forall (x : A),
 | 
				
			||||||
        {| x |} ∪ {|x|} = {|x|}.
 | 
					        {|x|} ∪ {|x|} = {|x|}.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  End FSet.
 | 
					  End FSet.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Arguments E {_}.
 | 
					 | 
				
			||||||
  Arguments U {_} _ _.
 | 
					 | 
				
			||||||
  Arguments L {_} _.
 | 
					 | 
				
			||||||
  Arguments assoc {_} _ _ _.
 | 
					  Arguments assoc {_} _ _ _.
 | 
				
			||||||
  Arguments comm {_} _ _.
 | 
					  Arguments comm {_} _ _.
 | 
				
			||||||
  Arguments nl {_} _.
 | 
					  Arguments nl {_} _.
 | 
				
			||||||
  Arguments nr {_} _.
 | 
					  Arguments nr {_} _.
 | 
				
			||||||
  Arguments idem {_} _.
 | 
					  Arguments idem {_} _.
 | 
				
			||||||
  Notation "{| x |}" :=  (L x).
 | 
					 | 
				
			||||||
  Infix "∪" := U (at level 8, right associativity).
 | 
					 | 
				
			||||||
  Notation "∅" := E.
 | 
					 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Section FSet_induction.
 | 
					  Section FSet_induction.
 | 
				
			||||||
    Variable A: Type.
 | 
					    Variable A: Type.
 | 
				
			||||||
@@ -74,12 +69,11 @@ Module Export FSet.
 | 
				
			|||||||
             {struct x}
 | 
					             {struct x}
 | 
				
			||||||
      : P x
 | 
					      : P x
 | 
				
			||||||
      := (match x return _ -> _ -> _ -> _ -> _ -> P x with
 | 
					      := (match x return _ -> _ -> _ -> _ -> _ -> P x with
 | 
				
			||||||
          | ∅ => fun _ _ _ _ _ => eP
 | 
					          | E => fun _ _ _ _ _ => eP
 | 
				
			||||||
          | {|a|} => fun _ _ _ _ _ => lP a
 | 
					          | L a => fun _ _ _ _ _ => lP a
 | 
				
			||||||
          | y ∪ z => fun _ _ _ _ _ => uP y z (FSet_ind y) (FSet_ind z)
 | 
					          | U y z => fun _ _ _ _ _ => uP y z (FSet_ind y) (FSet_ind z)
 | 
				
			||||||
          end) assocP commP nlP nrP idemP.
 | 
					          end) assocP commP nlP nrP idemP.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					 | 
				
			||||||
    Axiom FSet_ind_beta_assoc : forall (x y z : FSet A),
 | 
					    Axiom FSet_ind_beta_assoc : forall (x y z : FSet A),
 | 
				
			||||||
        apD FSet_ind (assoc x y z) =
 | 
					        apD FSet_ind (assoc x y z) =
 | 
				
			||||||
        (assocP x y z (FSet_ind x) (FSet_ind y) (FSet_ind z)).
 | 
					        (assocP x y z (FSet_ind x) (FSet_ind y) (FSet_ind z)).
 | 
				
			||||||
@@ -192,10 +186,6 @@ Module Export FSet.
 | 
				
			|||||||
 | 
					
 | 
				
			||||||
End FSet.
 | 
					End FSet.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
Notation "{| x |}" :=  (L x).
 | 
					 | 
				
			||||||
Infix "∪" := U (at level 8, right associativity).
 | 
					 | 
				
			||||||
Notation "∅" := E.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
Section set_sphere.
 | 
					Section set_sphere.
 | 
				
			||||||
  From HoTT.HIT Require Import Circle.
 | 
					  From HoTT.HIT Require Import Circle.
 | 
				
			||||||
  Context `{Univalence}.
 | 
					  Context `{Univalence}.
 | 
				
			||||||
@@ -274,10 +264,10 @@ Section set_sphere.
 | 
				
			|||||||
    - simpl. reflexivity.
 | 
					    - simpl. reflexivity.
 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Lemma FSet_S_ap : (nl (@E A)) = (nr E) -> idpath = loop.
 | 
					  Lemma FSet_S_ap : (nl (E A)) = (nr ∅) -> idpath = loop.
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    intros H1.
 | 
					    intros H1.
 | 
				
			||||||
    enough (ap FSet_to_S (nl E) = ap FSet_to_S (nr E)) as H'.
 | 
					    enough (ap FSet_to_S (nl ∅) = ap FSet_to_S (nr ∅)) as H'.
 | 
				
			||||||
    - rewrite FSet_rec_beta_nl in H'. 
 | 
					    - rewrite FSet_rec_beta_nl in H'. 
 | 
				
			||||||
      rewrite FSet_rec_beta_nr in H'.
 | 
					      rewrite FSet_rec_beta_nr in H'.
 | 
				
			||||||
      simpl in H'. unfold S1op_nr in H'.
 | 
					      simpl in H'. unfold S1op_nr in H'.
 | 
				
			||||||
 
 | 
				
			|||||||
@@ -1,35 +1,35 @@
 | 
				
			|||||||
(* Definition of Finite Sets as via cons lists *)
 | 
					(* Definition of Finite Sets as via cons lists *)
 | 
				
			||||||
Require Import HoTT HitTactics.
 | 
					Require Import HoTT HitTactics.
 | 
				
			||||||
 | 
					Require Export notation.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
Module Export FSetC.
 | 
					Module Export FSetC.
 | 
				
			||||||
  
 | 
					  
 | 
				
			||||||
  Section FSetC.
 | 
					  Section FSetC.
 | 
				
			||||||
 | 
					    Private Inductive FSetC (A : Type) : Type :=
 | 
				
			||||||
 | 
					    | Nil : FSetC A
 | 
				
			||||||
 | 
					    | Cns : A ->  FSetC A -> FSetC A.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					    Global Instance fset_empty : hasEmpty FSetC := Nil.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
    Variable A : Type.
 | 
					    Variable A : Type.
 | 
				
			||||||
 | 
					    Arguments Cns {_} _ _.
 | 
				
			||||||
    Private Inductive FSetC : Type :=
 | 
					 | 
				
			||||||
    | Nil : FSetC
 | 
					 | 
				
			||||||
    | Cns : A ->  FSetC -> FSetC.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
    Infix ";;" := Cns (at level 8, right associativity).
 | 
					    Infix ";;" := Cns (at level 8, right associativity).
 | 
				
			||||||
    Notation "∅" := Nil.
 | 
					 | 
				
			||||||
    
 | 
					    
 | 
				
			||||||
    Axiom dupl : forall (a: A) (x: FSetC),
 | 
					    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),
 | 
					    Axiom comm : forall (a b : A) (x : FSetC A),
 | 
				
			||||||
        a ;; b ;; x = b ;; a ;; x.
 | 
					        a ;; b ;; x = b ;; a ;; x.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
    Axiom trunc : IsHSet FSetC.
 | 
					    Axiom trunc : IsHSet (FSetC A).
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  End FSetC.
 | 
					  End FSetC.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Arguments Nil {_}.
 | 
					 | 
				
			||||||
  Arguments Cns {_} _ _.
 | 
					  Arguments Cns {_} _ _.
 | 
				
			||||||
  Arguments dupl {_} _ _.
 | 
					  Arguments dupl {_} _ _.
 | 
				
			||||||
  Arguments comm {_} _ _ _.
 | 
					  Arguments comm {_} _ _ _.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Infix ";;" := Cns (at level 8, right associativity).
 | 
					  Infix ";;" := Cns (at level 8, right associativity).
 | 
				
			||||||
  Notation "∅" := Nil.
 | 
					 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Section FSetC_induction.
 | 
					  Section FSetC_induction.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
@@ -84,7 +84,6 @@ Module Export FSetC.
 | 
				
			|||||||
      - apply commP. 
 | 
					      - apply commP. 
 | 
				
			||||||
    Defined.
 | 
					    Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					 | 
				
			||||||
    Definition FSetC_rec_beta_dupl : forall (a: A) (x : FSetC A),
 | 
					    Definition FSetC_rec_beta_dupl : forall (a: A) (x : FSetC A),
 | 
				
			||||||
        ap FSetC_rec (dupl a x) = duplP a (FSetC_rec x).
 | 
					        ap FSetC_rec (dupl a x) = duplP a (FSetC_rec x).
 | 
				
			||||||
    Proof.
 | 
					    Proof.
 | 
				
			||||||
@@ -106,11 +105,12 @@ Module Export FSetC.
 | 
				
			|||||||
  End FSetC_recursion.
 | 
					  End FSetC_recursion.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Instance FSetC_recursion A : HitRecursion (FSetC A) := {
 | 
					  Instance FSetC_recursion A : HitRecursion (FSetC A) :=
 | 
				
			||||||
                                                          indTy := _; recTy := _; 
 | 
					    {
 | 
				
			||||||
                                                          H_inductor := FSetC_ind A; H_recursor := FSetC_rec A }.
 | 
					      indTy := _; recTy := _; 
 | 
				
			||||||
 | 
					      H_inductor := FSetC_ind A; H_recursor := FSetC_rec A
 | 
				
			||||||
 | 
					    }.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
End FSetC.
 | 
					End FSetC.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
Infix ";;" := Cns (at level 8, right associativity).
 | 
					Infix ";;" := Cns (at level 8, right associativity).
 | 
				
			||||||
Notation "∅" := Nil.
 | 
					 | 
				
			||||||
@@ -1,50 +1,44 @@
 | 
				
			|||||||
(* Definitions of the Kuratowski-finite sets via a HIT *)
 | 
					(* Definitions of the Kuratowski-finite sets via a HIT *)
 | 
				
			||||||
Require Import HoTT.
 | 
					Require Import HoTT HitTactics.
 | 
				
			||||||
Require Import HitTactics.
 | 
					Require Export notation.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
Module Export FSet.
 | 
					Module Export FSet.
 | 
				
			||||||
  Section FSet.
 | 
					  Section FSet.
 | 
				
			||||||
 | 
					    Private Inductive FSet (A : Type) : Type :=
 | 
				
			||||||
 | 
					    | E : FSet A
 | 
				
			||||||
 | 
					    | L : A -> FSet A
 | 
				
			||||||
 | 
					    | U : FSet A -> FSet A -> FSet A.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					    Global Instance fset_empty : hasEmpty FSet := E.
 | 
				
			||||||
 | 
					    Global Instance fset_singleton : hasSingleton FSet := L.
 | 
				
			||||||
 | 
					    Global Instance fset_union : hasUnion FSet := U.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
    Variable A : Type.
 | 
					    Variable A : Type.
 | 
				
			||||||
    
 | 
					    
 | 
				
			||||||
    Private Inductive FSet : Type :=
 | 
					    Axiom assoc : forall (x y z : FSet A),
 | 
				
			||||||
    | E : FSet
 | 
					 | 
				
			||||||
    | L : A -> FSet
 | 
					 | 
				
			||||||
    | U : FSet -> FSet -> FSet.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
    Notation "{| x |}" :=  (L x).
 | 
					 | 
				
			||||||
    Infix "∪" := U (at level 8, right associativity).
 | 
					 | 
				
			||||||
    Notation "∅" := E.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
    Axiom assoc : forall (x y z : FSet ),
 | 
					 | 
				
			||||||
        x ∪ (y ∪ z) = (x ∪ y) ∪ z.
 | 
					        x ∪ (y ∪ z) = (x ∪ y) ∪ z.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
    Axiom comm : forall (x y : FSet),
 | 
					    Axiom comm : forall (x y : FSet A),
 | 
				
			||||||
        x ∪ y = y ∪ x.
 | 
					        x ∪ y = y ∪ x.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
    Axiom nl : forall (x : FSet),
 | 
					    Axiom nl : forall (x : FSet A),
 | 
				
			||||||
        ∅ ∪ x = x.
 | 
					        ∅ ∪ x = x.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
    Axiom nr : forall (x : FSet),
 | 
					    Axiom nr : forall (x : FSet A),
 | 
				
			||||||
        x ∪ ∅ = x.
 | 
					        x ∪ ∅ = x.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
    Axiom idem : forall (x : A),
 | 
					    Axiom idem : forall (x : A),
 | 
				
			||||||
        {| x |} ∪ {|x|} = {|x|}.
 | 
					        {|x|} ∪ {|x|} = {|x|}.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
    Axiom trunc : IsHSet FSet.
 | 
					    Axiom trunc : IsHSet (FSet A).
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  End FSet.
 | 
					  End FSet.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Arguments E {_}.
 | 
					 | 
				
			||||||
  Arguments U {_} _ _.
 | 
					 | 
				
			||||||
  Arguments L {_} _.
 | 
					 | 
				
			||||||
  Arguments assoc {_} _ _ _.
 | 
					  Arguments assoc {_} _ _ _.
 | 
				
			||||||
  Arguments comm {_} _ _.
 | 
					  Arguments comm {_} _ _.
 | 
				
			||||||
  Arguments nl {_} _.
 | 
					  Arguments nl {_} _.
 | 
				
			||||||
  Arguments nr {_} _.
 | 
					  Arguments nr {_} _.
 | 
				
			||||||
  Arguments idem {_} _.  
 | 
					  Arguments idem {_} _.  
 | 
				
			||||||
  Notation "{| x |}" :=  (L x).
 | 
					 | 
				
			||||||
  Infix "∪" := U (at level 8, right associativity).
 | 
					 | 
				
			||||||
  Notation "∅" := E.
 | 
					 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Section FSet_induction.
 | 
					  Section FSet_induction.
 | 
				
			||||||
    Variable A: Type.
 | 
					    Variable A: Type.
 | 
				
			||||||
@@ -194,10 +188,6 @@ Module Export FSet.
 | 
				
			|||||||
 | 
					
 | 
				
			||||||
End FSet.
 | 
					End FSet.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
Notation "{| x |}" :=  (L x).
 | 
					 | 
				
			||||||
Infix "∪" := U (at level 8, right associativity).
 | 
					 | 
				
			||||||
Notation "∅" := E.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
Lemma union_idem {A : Type} : forall x: FSet A, x ∪ x = x.
 | 
					Lemma union_idem {A : Type} : forall x: FSet A, x ∪ x = x.
 | 
				
			||||||
Proof.
 | 
					Proof.
 | 
				
			||||||
  hinduction ; try (intros ; apply set_path2).
 | 
					  hinduction ; try (intros ; apply set_path2).
 | 
				
			||||||
 
 | 
				
			|||||||
		Reference in New Issue
	
	Block a user