mirror of
				https://github.com/nmvdw/HITs-Examples
				synced 2025-11-03 23:23:51 +01:00 
			
		
		
		
	Some cleaning in notation
This commit is contained in:
		@@ -8,7 +8,7 @@ Section ext.
 | 
			
		||||
  Context `{Univalence}.
 | 
			
		||||
 | 
			
		||||
  Lemma subset_union_equiv
 | 
			
		||||
    : forall X Y : FSet A, subset X Y <~> U X Y = Y.
 | 
			
		||||
    : forall X Y : FSet A, X ⊆ Y <~> X ∪ Y = Y.
 | 
			
		||||
  Proof.
 | 
			
		||||
    intros X Y.
 | 
			
		||||
    eapply equiv_iff_hprop_uncurried.
 | 
			
		||||
@@ -20,8 +20,8 @@ Section ext.
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
  Lemma subset_isIn (X Y : FSet A) :
 | 
			
		||||
    (forall (a : A), isIn a X -> isIn a Y)
 | 
			
		||||
      <~> (subset X Y).
 | 
			
		||||
    (forall (a : A), a ∈ X -> a ∈ Y)
 | 
			
		||||
      <~> (X ⊆ Y).
 | 
			
		||||
  Proof.
 | 
			
		||||
    eapply equiv_iff_hprop_uncurried.
 | 
			
		||||
    split.
 | 
			
		||||
@@ -32,19 +32,17 @@ Section ext.
 | 
			
		||||
        apply f.
 | 
			
		||||
        apply tr ; reflexivity.
 | 
			
		||||
      + intros X1 X2 H1 H2 f.
 | 
			
		||||
        enough (subset X1 Y).
 | 
			
		||||
        enough (subset X2 Y).
 | 
			
		||||
        enough (X1 ⊆ Y).
 | 
			
		||||
        enough (X2 ⊆ Y).
 | 
			
		||||
        { split. apply X. apply X0. }
 | 
			
		||||
        * apply H2.
 | 
			
		||||
          intros a Ha.
 | 
			
		||||
          apply f.
 | 
			
		||||
          apply tr.
 | 
			
		||||
          refine (f _ (tr _)).
 | 
			
		||||
          right.
 | 
			
		||||
          apply Ha.
 | 
			
		||||
        * apply H1.
 | 
			
		||||
          intros a Ha.
 | 
			
		||||
          apply f.
 | 
			
		||||
          apply tr.
 | 
			
		||||
          refine (f _ (tr _)).
 | 
			
		||||
          left.
 | 
			
		||||
          apply Ha.
 | 
			
		||||
    - hinduction X ;
 | 
			
		||||
@@ -64,7 +62,7 @@ Section ext.
 | 
			
		||||
 | 
			
		||||
  (** ** Extensionality proof *)
 | 
			
		||||
 | 
			
		||||
  Lemma eq_subset' (X Y : FSet A) : X = Y <~> (U Y X = X) * (U X Y = Y).
 | 
			
		||||
  Lemma eq_subset' (X Y : FSet A) : X = Y <~> (Y ∪ X = X) * (X ∪ Y = Y).
 | 
			
		||||
  Proof.
 | 
			
		||||
    unshelve eapply BuildEquiv.
 | 
			
		||||
    { intro H'. rewrite H'. split; apply union_idem. }
 | 
			
		||||
@@ -76,20 +74,20 @@ Section ext.
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
  Lemma eq_subset (X Y : FSet A) :
 | 
			
		||||
    X = Y <~> (subset Y X * subset X Y).
 | 
			
		||||
    X = Y <~> (Y ⊆ X * X ⊆ Y).
 | 
			
		||||
  Proof.
 | 
			
		||||
    transitivity ((U Y X = X) * (U X Y = Y)).
 | 
			
		||||
    transitivity ((Y ∪ X = X) * (X ∪ Y = Y)).
 | 
			
		||||
    apply eq_subset'.
 | 
			
		||||
    symmetry.
 | 
			
		||||
    eapply equiv_functor_prod'; apply subset_union_equiv.
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
  Theorem fset_ext (X Y : FSet A) :
 | 
			
		||||
    X = Y <~> (forall (a : A), isIn a X = isIn a Y).
 | 
			
		||||
    X = Y <~> (forall (a : A), a ∈ X = a ∈ Y).
 | 
			
		||||
  Proof.
 | 
			
		||||
    refine (@equiv_compose' _ _ _ _ _) ; [ | apply eq_subset ].
 | 
			
		||||
    refine (@equiv_compose' _ ((forall a, isIn a Y -> isIn a X)
 | 
			
		||||
                               *(forall a, isIn a X -> isIn a Y)) _ _ _).
 | 
			
		||||
    refine (@equiv_compose' _ ((forall a, a ∈ Y -> a ∈ X)
 | 
			
		||||
                               *(forall a, a ∈ X -> a ∈ Y)) _ _ _).
 | 
			
		||||
    - apply equiv_iff_hprop_uncurried.
 | 
			
		||||
      split.
 | 
			
		||||
      * intros [H1 H2 a].
 | 
			
		||||
 
 | 
			
		||||
@@ -11,7 +11,7 @@ Section Length.
 | 
			
		||||
 | 
			
		||||
  Opaque isIn_b.
 | 
			
		||||
 | 
			
		||||
  Definition length (x: FSetC A) : nat.
 | 
			
		||||
  Definition length (x : FSetC A) : nat.
 | 
			
		||||
  Proof.
 | 
			
		||||
    simple refine (FSetC_ind A _ _ _ _ _ _ x ); simpl.
 | 
			
		||||
    - exact 0.
 | 
			
		||||
@@ -31,7 +31,7 @@ Section Length.
 | 
			
		||||
 | 
			
		||||
  Definition length_FSet (x: FSet A) := length (FSet_to_FSetC x).
 | 
			
		||||
 | 
			
		||||
  Lemma length_singleton: forall (a: A), length_FSet (L a) = 1.
 | 
			
		||||
  Lemma length_singleton: forall (a: A), length_FSet ({|a|}) = 1.
 | 
			
		||||
  Proof. 
 | 
			
		||||
    intro a.
 | 
			
		||||
    cbn. reflexivity. 
 | 
			
		||||
 
 | 
			
		||||
@@ -3,84 +3,78 @@ Require Import HoTT HitTactics.
 | 
			
		||||
Require Import representations.definition disjunction lattice.
 | 
			
		||||
 | 
			
		||||
Section operations.
 | 
			
		||||
Context {A : Type}.
 | 
			
		||||
Context `{Univalence}.
 | 
			
		||||
  Context {A : Type}.
 | 
			
		||||
  Context `{Univalence}.
 | 
			
		||||
 | 
			
		||||
Definition isIn : A -> FSet A -> hProp.
 | 
			
		||||
Proof.
 | 
			
		||||
intros a.
 | 
			
		||||
hrecursion.
 | 
			
		||||
- exists Empty.
 | 
			
		||||
  exact _.
 | 
			
		||||
- intro a'.
 | 
			
		||||
  exists (Trunc (-1) (a = a')).
 | 
			
		||||
  exact _.
 | 
			
		||||
- apply lor.
 | 
			
		||||
- intros ; symmetry ; apply lor_assoc.
 | 
			
		||||
- apply lor_commutative.
 | 
			
		||||
- apply lor_nl.
 | 
			
		||||
- apply lor_nr.
 | 
			
		||||
- intros ; apply lor_idem.
 | 
			
		||||
Defined.
 | 
			
		||||
  Definition isIn : A -> FSet A -> hProp.
 | 
			
		||||
  Proof.
 | 
			
		||||
    intros a.
 | 
			
		||||
    hrecursion.
 | 
			
		||||
    - exists Empty.
 | 
			
		||||
      exact _.
 | 
			
		||||
    - intro a'.
 | 
			
		||||
      exists (Trunc (-1) (a = a')).
 | 
			
		||||
      exact _.
 | 
			
		||||
    - apply lor.
 | 
			
		||||
    - intros ; symmetry ; apply lor_assoc.
 | 
			
		||||
    - apply lor_commutative.
 | 
			
		||||
    - apply lor_nl.
 | 
			
		||||
    - apply lor_nr.
 | 
			
		||||
    - 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 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.
 | 
			
		||||
intros P.
 | 
			
		||||
hrecursion.
 | 
			
		||||
- apply E.
 | 
			
		||||
- intro a.
 | 
			
		||||
  refine (if (P a) then L a else E).
 | 
			
		||||
- apply U.
 | 
			
		||||
- apply assoc.
 | 
			
		||||
- apply comm.
 | 
			
		||||
- apply nl.
 | 
			
		||||
- apply nr.
 | 
			
		||||
- intros; simpl. 
 | 
			
		||||
  destruct (P x).
 | 
			
		||||
  + apply idem.
 | 
			
		||||
  + apply nl.
 | 
			
		||||
Defined.
 | 
			
		||||
  Definition comprehension : 
 | 
			
		||||
    (A -> Bool) -> FSet A -> FSet A.
 | 
			
		||||
  Proof.
 | 
			
		||||
    intros P.
 | 
			
		||||
    hrecursion.
 | 
			
		||||
    - apply ∅.
 | 
			
		||||
    - intro a.
 | 
			
		||||
      refine (if (P a) then {|a|} else ∅).
 | 
			
		||||
    - apply U.
 | 
			
		||||
    - apply assoc.
 | 
			
		||||
    - apply comm.
 | 
			
		||||
    - apply nl.
 | 
			
		||||
    - apply nr.
 | 
			
		||||
    - intros; simpl. 
 | 
			
		||||
      destruct (P x).
 | 
			
		||||
      + apply idem.
 | 
			
		||||
      + apply nl.
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
Definition isEmpty :
 | 
			
		||||
  FSet A -> Bool.
 | 
			
		||||
Proof.
 | 
			
		||||
  hrecursion.
 | 
			
		||||
  - apply true.
 | 
			
		||||
  - apply (fun _ => false).
 | 
			
		||||
  - apply andb.
 | 
			
		||||
  - intros. symmetry. eauto with lattice_hints typeclass_instances.
 | 
			
		||||
  - eauto with bool_lattice_hints typeclass_instances.
 | 
			
		||||
  - eauto with bool_lattice_hints typeclass_instances.
 | 
			
		||||
  - eauto with bool_lattice_hints typeclass_instances.
 | 
			
		||||
  - eauto with bool_lattice_hints typeclass_instances.
 | 
			
		||||
Defined.    
 | 
			
		||||
  Definition isEmpty :
 | 
			
		||||
    FSet A -> Bool.
 | 
			
		||||
  Proof.
 | 
			
		||||
    simple refine (FSet_rec _ _ _ true (fun _ => false) andb _ _ _ _ _)
 | 
			
		||||
    ; try eauto with bool_lattice_hints typeclass_instances.
 | 
			
		||||
    intros ; symmetry ; eauto with lattice_hints typeclass_instances.
 | 
			
		||||
  Defined.    
 | 
			
		||||
  
 | 
			
		||||
End operations.
 | 
			
		||||
 | 
			
		||||
 
 | 
			
		||||
@@ -7,7 +7,7 @@ Section operations_isIn.
 | 
			
		||||
  Context {A : Type}.
 | 
			
		||||
  Context `{Univalence}.
 | 
			
		||||
 | 
			
		||||
  Lemma union_idem : forall x: FSet A, U x x = x.
 | 
			
		||||
  Lemma union_idem : forall x: FSet A, x ∪ x = x.
 | 
			
		||||
  Proof.
 | 
			
		||||
    hinduction ; try (intros ; apply set_path2).
 | 
			
		||||
    - apply nl.
 | 
			
		||||
@@ -24,7 +24,7 @@ Section operations_isIn.
 | 
			
		||||
 | 
			
		||||
  (** ** Properties about subset relation. *)
 | 
			
		||||
  Lemma subset_union (X Y : FSet A) : 
 | 
			
		||||
    subset X Y -> U X Y = Y.
 | 
			
		||||
    X ⊆ Y -> X ∪ Y = Y.
 | 
			
		||||
  Proof.
 | 
			
		||||
    hinduction X ; try (intros; apply path_forall; intro; apply set_path2).
 | 
			
		||||
    - intros. apply nl.
 | 
			
		||||
@@ -54,7 +54,7 @@ Section operations_isIn.
 | 
			
		||||
  Defined.
 | 
			
		||||
  
 | 
			
		||||
  Lemma subset_union_l (X : FSet A) :
 | 
			
		||||
    forall Y, subset X (U X Y).
 | 
			
		||||
    forall Y, X ⊆ X ∪ Y.
 | 
			
		||||
  Proof.
 | 
			
		||||
    hinduction X ; try (repeat (intro; intros; apply path_forall);
 | 
			
		||||
                        intro ; apply equiv_hprop_allpath ; apply _).
 | 
			
		||||
@@ -69,8 +69,8 @@ Section operations_isIn.
 | 
			
		||||
 | 
			
		||||
  (* simplify it using extensionality *)
 | 
			
		||||
  Lemma comprehension_or : forall ϕ ψ (x: FSet A),
 | 
			
		||||
      comprehension (fun a => orb (ϕ a) (ψ a)) x = U (comprehension ϕ x) 
 | 
			
		||||
                                                     (comprehension ψ x).
 | 
			
		||||
      comprehension (fun a => orb (ϕ a) (ψ a)) x
 | 
			
		||||
      = (comprehension ϕ x) ∪ (comprehension ψ x).
 | 
			
		||||
  Proof.
 | 
			
		||||
    intros ϕ ψ.
 | 
			
		||||
    hinduction ; try (intros; apply set_path2). 
 | 
			
		||||
@@ -101,15 +101,15 @@ Section properties.
 | 
			
		||||
  Context `{Univalence}.
 | 
			
		||||
 | 
			
		||||
  (** isIn properties *)
 | 
			
		||||
  Definition empty_isIn (a: A) : isIn a E -> Empty := idmap.
 | 
			
		||||
  Definition empty_isIn (a: A) : a ∈ E -> Empty := idmap.
 | 
			
		||||
  
 | 
			
		||||
  Definition singleton_isIn (a b: A) : isIn a (L 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)
 | 
			
		||||
    : isIn a (U X Y) = isIn a X ∨ isIn a Y := idpath.
 | 
			
		||||
    : a ∈ X ∪ Y = a ∈ X ∨ a ∈ Y := idpath.
 | 
			
		||||
 | 
			
		||||
  Lemma comprehension_isIn (ϕ : A -> Bool) (a : A) : forall X : FSet A,
 | 
			
		||||
      isIn a (comprehension ϕ X) = if ϕ a then isIn a X else False_hp.
 | 
			
		||||
      a ∈ (comprehension ϕ X) = if ϕ a then a ∈ X else False_hp.
 | 
			
		||||
  Proof.
 | 
			
		||||
    hinduction ; try (intros ; apply set_path2) ; cbn.
 | 
			
		||||
    - destruct (ϕ a) ; reflexivity.
 | 
			
		||||
@@ -139,7 +139,7 @@ Section properties.
 | 
			
		||||
 | 
			
		||||
  (* The proof can be simplified using extensionality *)
 | 
			
		||||
  (** comprehension properties *)
 | 
			
		||||
  Lemma comprehension_false Y : comprehension (fun (_ : A) => false) Y = E.
 | 
			
		||||
  Lemma comprehension_false Y : comprehension (fun (_ : A) => false) Y = ∅.
 | 
			
		||||
  Proof.
 | 
			
		||||
    hrecursion Y; try (intros; apply set_path2).
 | 
			
		||||
    - reflexivity.
 | 
			
		||||
@@ -151,7 +151,7 @@ Section properties.
 | 
			
		||||
 | 
			
		||||
  (* Can be simplified using extensionality *)
 | 
			
		||||
  Lemma comprehension_subset : forall ϕ (X : FSet A),
 | 
			
		||||
      U (comprehension ϕ X) X = X.
 | 
			
		||||
      (comprehension ϕ X) ∪ X = X.
 | 
			
		||||
  Proof.
 | 
			
		||||
    intros ϕ.
 | 
			
		||||
    hrecursion; try (intros ; apply set_path2) ; cbn.
 | 
			
		||||
@@ -171,7 +171,7 @@ Section properties.
 | 
			
		||||
      reflexivity.
 | 
			
		||||
  Defined.
 | 
			
		||||
  
 | 
			
		||||
  Lemma merely_choice : forall X : FSet A, hor (X = E) (hexists (fun a => isIn a X)).
 | 
			
		||||
  Lemma merely_choice : forall X : FSet A, hor (X = ∅) (hexists (fun a => a ∈ X)).
 | 
			
		||||
  Proof.
 | 
			
		||||
    hinduction; try (intros; apply equiv_hprop_allpath ; apply _).
 | 
			
		||||
    - apply (tr (inl idpath)).
 | 
			
		||||
 
 | 
			
		||||
@@ -29,7 +29,7 @@ Section operations_isIn.
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
  Lemma L_isIn (a b : A) :
 | 
			
		||||
    isIn a {|b|} -> a = b.
 | 
			
		||||
    a ∈ {|b|} -> a = b.
 | 
			
		||||
  Proof.
 | 
			
		||||
    intros. strip_truncations. assumption.
 | 
			
		||||
  Defined.
 | 
			
		||||
@@ -63,7 +63,7 @@ Section operations_isIn.
 | 
			
		||||
  
 | 
			
		||||
  (* Union and membership *)
 | 
			
		||||
  Lemma union_isIn_b (X Y : FSet A) (a : A) :
 | 
			
		||||
    isIn_b a (U X Y) = orb (isIn_b a X) (isIn_b a Y).
 | 
			
		||||
    isIn_b a (X ∪ Y) = orb (isIn_b a X) (isIn_b a Y).
 | 
			
		||||
  Proof.
 | 
			
		||||
    unfold isIn_b ; unfold dec.
 | 
			
		||||
    simpl.
 | 
			
		||||
@@ -109,7 +109,7 @@ Section SetLattice.
 | 
			
		||||
 | 
			
		||||
  Instance fset_max : maximum (FSet A) := U.
 | 
			
		||||
  Instance fset_min : minimum (FSet A) := intersection.
 | 
			
		||||
  Instance fset_bot : bottom (FSet A) := E.
 | 
			
		||||
  Instance fset_bot : bottom (FSet A) := ∅.
 | 
			
		||||
  
 | 
			
		||||
  Instance lattice_fset : Lattice (FSet A).
 | 
			
		||||
  Proof.
 | 
			
		||||
@@ -133,7 +133,7 @@ Section comprehension_properties.
 | 
			
		||||
  Defined.
 | 
			
		||||
  
 | 
			
		||||
  (** comprehension properties *)
 | 
			
		||||
  Lemma comprehension_false Y : comprehension (fun (_ : A) => false) Y = E.
 | 
			
		||||
  Lemma comprehension_false Y : comprehension (fun (_ : A) => false) Y = ∅.
 | 
			
		||||
  Proof.
 | 
			
		||||
    toBool.
 | 
			
		||||
  Defined.
 | 
			
		||||
@@ -145,7 +145,7 @@ Section comprehension_properties.
 | 
			
		||||
  Defined.
 | 
			
		||||
  
 | 
			
		||||
  Lemma comprehension_subset : forall ϕ (X : FSet A),
 | 
			
		||||
      U (comprehension ϕ X) X = X.
 | 
			
		||||
      (comprehension ϕ X) ∪ X = X.
 | 
			
		||||
  Proof.
 | 
			
		||||
    toBool.
 | 
			
		||||
  Defined.
 | 
			
		||||
@@ -159,7 +159,7 @@ Section dec_eq.
 | 
			
		||||
  Instance fsets_dec_eq : DecidablePaths (FSet A).
 | 
			
		||||
  Proof.
 | 
			
		||||
    intros X Y.
 | 
			
		||||
    apply (decidable_equiv' ((subset Y X) * (subset X Y)) (eq_subset X Y)^-1).
 | 
			
		||||
    apply (decidable_equiv' ((Y ⊆ X) * (X ⊆ Y)) (eq_subset X Y)^-1).
 | 
			
		||||
    apply decidable_prod.
 | 
			
		||||
  Defined.
 | 
			
		||||
 | 
			
		||||
 
 | 
			
		||||
		Reference in New Issue
	
	Block a user