mirror of
				https://github.com/nmvdw/HITs-Examples
				synced 2025-11-04 07:33:51 +01:00 
			
		
		
		
	Improved notatio
This commit is contained in:
		@@ -12,25 +12,35 @@ Section Iso.
 | 
			
		||||
  Proof.
 | 
			
		||||
    hrecursion.
 | 
			
		||||
    - apply E.
 | 
			
		||||
    - intros a x. apply (U (L a) x).
 | 
			
		||||
    - intros a x.
 | 
			
		||||
      apply ({|a|} ∪ x).
 | 
			
		||||
    - intros. cbn.  
 | 
			
		||||
      etransitivity. apply assoc.
 | 
			
		||||
      apply (ap (fun y => U y x)).
 | 
			
		||||
      apply (ap (∪ x)).
 | 
			
		||||
      apply idem.
 | 
			
		||||
    - intros. cbn.
 | 
			
		||||
      etransitivity. apply assoc.
 | 
			
		||||
      etransitivity. refine (ap (fun y => U y x) _ ).
 | 
			
		||||
      etransitivity. refine (ap (∪ x) _ ).
 | 
			
		||||
      apply FSet.comm.
 | 
			
		||||
      symmetry. 
 | 
			
		||||
      apply assoc.
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
  Definition FSet_to_FSetC: FSet A -> FSetC A :=
 | 
			
		||||
    FSet_rec A (FSetC A) (FSetC.trunc A) Nil singleton append append_assoc 
 | 
			
		||||
             append_comm append_nl append_nr singleton_idem.
 | 
			
		||||
  Definition FSet_to_FSetC: FSet A -> FSetC A.
 | 
			
		||||
  Proof.
 | 
			
		||||
    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), 
 | 
			
		||||
      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.
 | 
			
		||||
    intros x. 
 | 
			
		||||
    hrecursion x; try (intros; apply path_forall; intro; apply set_path2).
 | 
			
		||||
 
 | 
			
		||||
@@ -8,7 +8,7 @@ Proof.
 | 
			
		||||
  hrecursion.
 | 
			
		||||
  - exact ∅.
 | 
			
		||||
  - intro a. exact {| f a |}.
 | 
			
		||||
  - exact U.
 | 
			
		||||
  - intros X Y. apply (X ∪ Y).
 | 
			
		||||
  - apply assoc.
 | 
			
		||||
  - apply comm.
 | 
			
		||||
  - apply nl.
 | 
			
		||||
@@ -39,7 +39,7 @@ Proof.
 | 
			
		||||
hrecursion.
 | 
			
		||||
- exact ∅.
 | 
			
		||||
- exact idmap.
 | 
			
		||||
- exact U.
 | 
			
		||||
- intros X Y. apply (X ∪ Y).
 | 
			
		||||
- apply assoc.
 | 
			
		||||
- apply comm.
 | 
			
		||||
- apply nl.
 | 
			
		||||
 
 | 
			
		||||
@@ -23,32 +23,6 @@ Section operations.
 | 
			
		||||
    - intros ; apply lor_idem.
 | 
			
		||||
  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 : 
 | 
			
		||||
    (A -> Bool) -> FSet A -> FSet A.
 | 
			
		||||
  Proof.
 | 
			
		||||
@@ -57,7 +31,7 @@ Section operations.
 | 
			
		||||
    - apply ∅.
 | 
			
		||||
    - intro a.
 | 
			
		||||
      refine (if (P a) then {|a|} else ∅).
 | 
			
		||||
    - apply U.
 | 
			
		||||
    - apply (∪).
 | 
			
		||||
    - apply assoc.
 | 
			
		||||
    - apply comm.
 | 
			
		||||
    - apply nl.
 | 
			
		||||
@@ -82,7 +56,7 @@ Section operations.
 | 
			
		||||
    - apply ∅.
 | 
			
		||||
    - intro b.
 | 
			
		||||
      apply {|(a, b)|}.
 | 
			
		||||
    - apply U.
 | 
			
		||||
    - apply (∪).
 | 
			
		||||
    - intros X Y Z ; apply assoc.
 | 
			
		||||
    - intros X Y ; apply comm.
 | 
			
		||||
    - intros ; apply nl.
 | 
			
		||||
@@ -97,7 +71,7 @@ Section operations.
 | 
			
		||||
    - apply ∅.
 | 
			
		||||
    - intro a.
 | 
			
		||||
      apply (single_product a Y).
 | 
			
		||||
    - apply U.
 | 
			
		||||
    - apply (∪).
 | 
			
		||||
    - intros ; apply assoc.
 | 
			
		||||
    - intros ; apply comm.
 | 
			
		||||
    - intros ; apply nl.
 | 
			
		||||
@@ -107,5 +81,48 @@ Section 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).
 | 
			
		||||
@@ -4,9 +4,9 @@ Require Import representations.cons_repr.
 | 
			
		||||
 | 
			
		||||
Section operations.
 | 
			
		||||
 | 
			
		||||
  Context {A : Type}.
 | 
			
		||||
 | 
			
		||||
  Definition append  (x y: FSetC A) : FSetC A.
 | 
			
		||||
  Global Instance fsetc_union : hasUnion FSetC.
 | 
			
		||||
  Proof.
 | 
			
		||||
    intros A x y.
 | 
			
		||||
    hinduction x.
 | 
			
		||||
    - apply y.
 | 
			
		||||
    - apply Cns.
 | 
			
		||||
@@ -14,6 +14,6 @@ Section operations.
 | 
			
		||||
    - apply comm.
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
  Definition singleton (a:A) : FSetC A := a;;∅.
 | 
			
		||||
  Global Instance fsetc_singleton : hasSingleton FSetC := fun A a => a;;∅.
 | 
			
		||||
 | 
			
		||||
End operations.
 | 
			
		||||
@@ -33,7 +33,7 @@ Section decidable_A.
 | 
			
		||||
  Proof.
 | 
			
		||||
    hinduction ; try (intros ; apply path_ishprop).
 | 
			
		||||
    - intro ; apply _.
 | 
			
		||||
    - intros ; apply _. 
 | 
			
		||||
    - intros. apply _. 
 | 
			
		||||
    - intros ; apply _.
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
 
 | 
			
		||||
@@ -9,17 +9,23 @@ Section characterize_isIn.
 | 
			
		||||
  Context `{Univalence}.
 | 
			
		||||
 | 
			
		||||
  (** 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 union_isIn (X Y : FSet A) (a : A)
 | 
			
		||||
    : a ∈ X ∪ Y = a ∈ X ∨ a ∈ Y := idpath.
 | 
			
		||||
 | 
			
		||||
  Lemma comprehension_isIn (ϕ : A -> Bool) (a : A) : forall X : FSet A,
 | 
			
		||||
      a ∈ (comprehension ϕ X) = if ϕ a then a ∈ X else False_hp.
 | 
			
		||||
  Lemma comprehension_union (ϕ : A -> Bool) : forall X Y : FSet A,
 | 
			
		||||
      {|X ∪ Y & ϕ|} = {|X & ϕ|} ∪ {|Y & ϕ|}.
 | 
			
		||||
  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.
 | 
			
		||||
    - intros b.
 | 
			
		||||
      assert (forall c d, ϕ a = c -> ϕ b = d ->
 | 
			
		||||
@@ -36,6 +42,8 @@ Section characterize_isIn.
 | 
			
		||||
      }
 | 
			
		||||
      apply (X (ϕ a) (ϕ b) idpath idpath).
 | 
			
		||||
    - intros X Y H1 H2.
 | 
			
		||||
      rewrite comprehension_union.
 | 
			
		||||
      rewrite union_isIn.
 | 
			
		||||
      rewrite H1, H2.
 | 
			
		||||
      destruct (ϕ a).
 | 
			
		||||
      * reflexivity.
 | 
			
		||||
@@ -44,11 +52,14 @@ Section characterize_isIn.
 | 
			
		||||
           destruct Z ; assumption.
 | 
			
		||||
        ** intros ; apply tr ; right ; assumption.
 | 
			
		||||
  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),
 | 
			
		||||
      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.
 | 
			
		||||
    hinduction ; try (intros ; apply path_ishprop).
 | 
			
		||||
    - apply path_hprop ; symmetry ; apply prod_empty_r.
 | 
			
		||||
@@ -62,16 +73,16 @@ Section characterize_isIn.
 | 
			
		||||
        rewrite Z1, Z2.
 | 
			
		||||
        apply (tr idpath).
 | 
			
		||||
    - intros X1 X2 HX1 HX2.
 | 
			
		||||
      apply path_iff_hprop.
 | 
			
		||||
      *  intros X.
 | 
			
		||||
         strip_truncations.
 | 
			
		||||
         destruct X as [H1 | H1] ; rewrite ?HX1, ?HX2 in H1 ; destruct H1 as [H1 H2].
 | 
			
		||||
         ** split.
 | 
			
		||||
            *** apply H1.
 | 
			
		||||
            *** apply (tr(inl H2)).
 | 
			
		||||
         ** split.
 | 
			
		||||
            *** apply H1.
 | 
			
		||||
            *** apply (tr(inr H2)).
 | 
			
		||||
      apply path_iff_hprop ; rewrite ?union_isIn.
 | 
			
		||||
      * intros X.
 | 
			
		||||
        strip_truncations.
 | 
			
		||||
        destruct X as [H1 | H1] ; rewrite ?HX1, ?HX2 in H1 ; destruct H1 as [H1 H2].
 | 
			
		||||
        ** split.
 | 
			
		||||
           *** apply H1.
 | 
			
		||||
           *** apply (tr(inl H2)).
 | 
			
		||||
        ** split.
 | 
			
		||||
           *** apply H1.
 | 
			
		||||
           *** apply (tr(inr H2)).
 | 
			
		||||
      * intros [H1 H2].
 | 
			
		||||
        strip_truncations.
 | 
			
		||||
        apply tr.
 | 
			
		||||
@@ -84,15 +95,16 @@ Section characterize_isIn.
 | 
			
		||||
  Defined.
 | 
			
		||||
  
 | 
			
		||||
  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.
 | 
			
		||||
    hinduction X ; try (intros ; apply path_ishprop).
 | 
			
		||||
    - apply path_hprop ; symmetry ; apply prod_empty_l.
 | 
			
		||||
    - intros.
 | 
			
		||||
      apply isIn_singleproduct.
 | 
			
		||||
    - intros X1 X2 HX1 HX2.
 | 
			
		||||
      rewrite (union_isIn (product X1 Y)).
 | 
			
		||||
      rewrite HX1, HX2.
 | 
			
		||||
      apply path_iff_hprop.
 | 
			
		||||
      apply path_iff_hprop ; rewrite ?union_isIn.
 | 
			
		||||
      * intros X.
 | 
			
		||||
        strip_truncations.
 | 
			
		||||
        destruct X as [[H3 H4] | [H3 H4]] ; split ; try (apply H4).
 | 
			
		||||
@@ -104,7 +116,7 @@ Section characterize_isIn.
 | 
			
		||||
        ** left ; split ; assumption.
 | 
			
		||||
        ** right ; split ; assumption.
 | 
			
		||||
  Defined.
 | 
			
		||||
End characterize_isIn.
 | 
			
		||||
End product_isIn.
 | 
			
		||||
 | 
			
		||||
Ltac simplify_isIn :=
 | 
			
		||||
  repeat (rewrite union_isIn
 | 
			
		||||
@@ -120,8 +132,17 @@ Section properties.
 | 
			
		||||
  Context {A : Type}.
 | 
			
		||||
  Context `{Univalence}.
 | 
			
		||||
 | 
			
		||||
  Instance: bottom(FSet A) := ∅.
 | 
			
		||||
  Instance: maximum(FSet A) := U.
 | 
			
		||||
  Instance: bottom(FSet A).
 | 
			
		||||
  Proof.
 | 
			
		||||
    unfold bottom.
 | 
			
		||||
    apply ∅.
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
  Instance: maximum(FSet A).
 | 
			
		||||
  Proof.
 | 
			
		||||
    intros x y.
 | 
			
		||||
    apply (x ∪ y).
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
  Global Instance joinsemilattice_fset : JoinSemiLattice (FSet A).
 | 
			
		||||
  Proof.
 | 
			
		||||
@@ -129,26 +150,32 @@ Section properties.
 | 
			
		||||
  Defined.  
 | 
			
		||||
 | 
			
		||||
  (** comprehension properties *)
 | 
			
		||||
  Lemma comprehension_false Y : comprehension (fun (_ : A) => false) Y = ∅.
 | 
			
		||||
  Lemma comprehension_false X : {|X & (fun (_ : A) => false)|} = ∅.
 | 
			
		||||
  Proof.
 | 
			
		||||
    toHProp.
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
  Lemma comprehension_subset : forall ϕ (X : FSet A),
 | 
			
		||||
      (comprehension ϕ X) ∪ X = X.
 | 
			
		||||
      {|X & ϕ|} ∪ X = X.
 | 
			
		||||
  Proof.
 | 
			
		||||
    toHProp.
 | 
			
		||||
    destruct (ϕ a) ; eauto with lattice_hints typeclass_instances.
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
  Lemma comprehension_or : forall ϕ ψ (x: FSet A),
 | 
			
		||||
      comprehension (fun a => orb (ϕ a) (ψ a)) x
 | 
			
		||||
      = (comprehension ϕ x) ∪ (comprehension ψ x).
 | 
			
		||||
  Lemma comprehension_or : forall ϕ ψ (X : FSet A),
 | 
			
		||||
      {|X & (fun a => orb (ϕ a) (ψ a))|}
 | 
			
		||||
      = {|X & ϕ|} ∪ {|X & ψ|}.
 | 
			
		||||
  Proof.
 | 
			
		||||
    toHProp.
 | 
			
		||||
    symmetry ; destruct (ϕ a) ; destruct (ψ a)
 | 
			
		||||
    ; eauto with lattice_hints typeclass_instances.
 | 
			
		||||
  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)).
 | 
			
		||||
  Proof.
 | 
			
		||||
@@ -161,7 +188,7 @@ Section properties.
 | 
			
		||||
      destruct TX as [XE | HX] ; destruct TY as [YE | HY] ; try(strip_truncations ; apply tr).
 | 
			
		||||
      * refine (tr (inl _)).
 | 
			
		||||
        rewrite XE, YE.
 | 
			
		||||
        apply (union_idem E).
 | 
			
		||||
        apply (union_idem ∅).
 | 
			
		||||
      * destruct HY as [a Ya].
 | 
			
		||||
        refine (inr (tr _)).
 | 
			
		||||
        exists a.
 | 
			
		||||
 
 | 
			
		||||
@@ -6,38 +6,30 @@ From fsets Require Import operations_cons_repr.
 | 
			
		||||
Section properties.
 | 
			
		||||
  Context {A : Type}.
 | 
			
		||||
 | 
			
		||||
  Lemma append_nl: 
 | 
			
		||||
    forall (x: FSetC A), append ∅ x = x. 
 | 
			
		||||
  Proof.
 | 
			
		||||
    intro. reflexivity.
 | 
			
		||||
  Defined.
 | 
			
		||||
  Definition append_nl : forall (x: FSetC A), ∅ ∪ x = x
 | 
			
		||||
    := fun _ => idpath. 
 | 
			
		||||
 | 
			
		||||
  Lemma append_nr: 
 | 
			
		||||
    forall (x: FSetC A), append x ∅ = x.
 | 
			
		||||
  Lemma append_nr : forall (x: FSetC A), x ∪ ∅ = x.
 | 
			
		||||
  Proof.
 | 
			
		||||
    hinduction; try (intros; apply set_path2).
 | 
			
		||||
    -  reflexivity.
 | 
			
		||||
    -  intros. apply (ap (fun y => a ;; y) X).
 | 
			
		||||
    -  intros. apply (ap (fun y => a;;y) X).
 | 
			
		||||
  Defined.
 | 
			
		||||
  
 | 
			
		||||
  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.
 | 
			
		||||
    intro x; hinduction x; try (intros; apply set_path2).
 | 
			
		||||
    hinduction
 | 
			
		||||
    ; try (intros ; apply path_forall ; intro
 | 
			
		||||
           ; apply path_forall ; intro ;  apply set_path2).
 | 
			
		||||
    - reflexivity.
 | 
			
		||||
    - intros a x HR y z. 
 | 
			
		||||
      specialize (HR y z).
 | 
			
		||||
      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.
 | 
			
		||||
      apply (ap (fun y => a;;y) HR). 
 | 
			
		||||
  Defined.
 | 
			
		||||
  
 | 
			
		||||
  Lemma append_singleton: forall (a: A) (x: FSetC A), 
 | 
			
		||||
      a ;; x = append x (a ;; ∅).
 | 
			
		||||
      a ;; x = x ∪ (a ;; ∅).
 | 
			
		||||
  Proof.
 | 
			
		||||
    intro a. hinduction; try (intros; apply set_path2).
 | 
			
		||||
    - reflexivity.
 | 
			
		||||
@@ -47,29 +39,26 @@ Section properties.
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
  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.
 | 
			
		||||
    intro x1. 
 | 
			
		||||
    hinduction x1;  try (intros; apply set_path2).
 | 
			
		||||
    hinduction ;  try (intros ; apply path_forall ; intro ; apply set_path2).
 | 
			
		||||
    - intros. symmetry. apply append_nr. 
 | 
			
		||||
    - intros a x1 HR x2.
 | 
			
		||||
      etransitivity.
 | 
			
		||||
      apply (ap (fun y => a;;y) (HR x2)).  
 | 
			
		||||
      transitivity  (append (append x2 x1) (a;;∅)).
 | 
			
		||||
      transitivity  ((x2 ∪ x1) ∪ (a;;∅)).
 | 
			
		||||
      + apply append_singleton. 
 | 
			
		||||
      + etransitivity.
 | 
			
		||||
    	* symmetry. apply append_assoc.
 | 
			
		||||
    	* simple refine (ap (fun y => append x2 y) (append_singleton _ _)^).
 | 
			
		||||
    - intros. apply path_forall.
 | 
			
		||||
      intro. apply set_path2.
 | 
			
		||||
    - intros. apply path_forall.
 | 
			
		||||
      intro. apply set_path2.  
 | 
			
		||||
    	* simple refine (ap (x2 ∪) (append_singleton _ _)^).
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
  Lemma singleton_idem: forall (a: A), 
 | 
			
		||||
      append (singleton a) (singleton a) = singleton a.
 | 
			
		||||
      {|a|} ∪ {|a|} = {|a|}.
 | 
			
		||||
  Proof.
 | 
			
		||||
    unfold singleton. intro. cbn. apply dupl.
 | 
			
		||||
    unfold singleton.
 | 
			
		||||
    intro.
 | 
			
		||||
    apply dupl.
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
End properties.
 | 
			
		||||
 
 | 
			
		||||
@@ -71,10 +71,10 @@ Section operations_isIn.
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
  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.
 | 
			
		||||
    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 (ϕ a) ; try reflexivity ; try contradiction.
 | 
			
		||||
  Defined.
 | 
			
		||||
@@ -105,10 +105,24 @@ Section SetLattice.
 | 
			
		||||
  Context {A_deceq : DecidablePaths A}.
 | 
			
		||||
  Context `{Univalence}.
 | 
			
		||||
 | 
			
		||||
  Instance fset_max : maximum (FSet A) := U.
 | 
			
		||||
  Instance fset_min : minimum (FSet A) := intersection.
 | 
			
		||||
  Instance fset_bot : bottom (FSet A) := ∅.
 | 
			
		||||
  Instance fset_max : maximum (FSet A).
 | 
			
		||||
  Proof.
 | 
			
		||||
    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).
 | 
			
		||||
  Proof.
 | 
			
		||||
    split; toBool.
 | 
			
		||||
@@ -116,40 +130,6 @@ Section 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 *)
 | 
			
		||||
Section dec_eq.
 | 
			
		||||
  Context (A : Type) `{DecidablePaths A} `{Univalence}.
 | 
			
		||||
 
 | 
			
		||||
		Reference in New Issue
	
	Block a user