mirror of
				https://github.com/nmvdw/HITs-Examples
				synced 2025-11-03 23:23:51 +01:00 
			
		
		
		
	Update the code to match the latest HoTT
HoTT commit 3526c344c47d32f5d4d268658031777239ec952b
This commit is contained in:
		@@ -51,7 +51,7 @@ Section ListToSet.
 | 
				
			|||||||
    induction l ; unfold member in * ; simpl in *.
 | 
					    induction l ; unfold member in * ; simpl in *.
 | 
				
			||||||
    - reflexivity.
 | 
					    - reflexivity.
 | 
				
			||||||
    - rewrite IHl.
 | 
					    - rewrite IHl.
 | 
				
			||||||
      unfold hor, merely, lor.
 | 
					      unfold hor, merely.
 | 
				
			||||||
      apply path_iff_hprop ; intros z ; strip_truncations ; destruct z as [z1 | z2].
 | 
					      apply path_iff_hprop ; intros z ; strip_truncations ; destruct z as [z1 | z2].
 | 
				
			||||||
      * apply (tr (inl (tr z1))).
 | 
					      * apply (tr (inl (tr z1))).
 | 
				
			||||||
      * apply (tr (inr z2)).
 | 
					      * apply (tr (inr z2)).
 | 
				
			||||||
 
 | 
				
			|||||||
@@ -1,290 +1,3 @@
 | 
				
			|||||||
(** Some examples of lattices. *)
 | 
					(** Some examples of lattices. *)
 | 
				
			||||||
Require Export HoTT lattice_interface.
 | 
					Require Export HoTT lattice_interface.
 | 
				
			||||||
 | 
					From HoTT.Classes.implementations Require Export hprop_lattice bool pointwise.
 | 
				
			||||||
(** [Bool] is a lattice. *)
 | 
					 | 
				
			||||||
Section BoolLattice.
 | 
					 | 
				
			||||||
  Ltac solve_bool :=
 | 
					 | 
				
			||||||
    let x := fresh in
 | 
					 | 
				
			||||||
    repeat (intro x ; destruct x)
 | 
					 | 
				
			||||||
    ; compute
 | 
					 | 
				
			||||||
    ; auto
 | 
					 | 
				
			||||||
    ; try contradiction.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Instance maximum_bool : Join Bool := orb.
 | 
					 | 
				
			||||||
  Instance minimum_bool : Meet Bool := andb.
 | 
					 | 
				
			||||||
  Instance bottom_bool : Bottom Bool := false.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Global Instance boundedjoinsemilattice_bool : BoundedJoinSemiLattice Bool.
 | 
					 | 
				
			||||||
  Proof. repeat (split ; (apply _ || solve_bool)). Defined.
 | 
					 | 
				
			||||||
  Global Instance meetsemilattice_bool : MeetSemiLattice Bool.
 | 
					 | 
				
			||||||
  Proof. repeat (split ; (apply _ || solve_bool)). Defined.
 | 
					 | 
				
			||||||
  Global Instance boundedmeetsemilattice_bool : @BoundedSemiLattice Bool (⊓) true.
 | 
					 | 
				
			||||||
  Proof. repeat (split ; (apply _ || solve_bool)). Defined.
 | 
					 | 
				
			||||||
  Global Instance lattice_bool : Lattice Bool.
 | 
					 | 
				
			||||||
  Proof. split ; (apply _ || solve_bool). Defined.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Definition and_true : forall b, andb b true = b.
 | 
					 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    solve_bool.
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Definition and_false : forall b, andb b false = false.
 | 
					 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    solve_bool.
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Definition dist₁ : forall b₁ b₂ b₃,
 | 
					 | 
				
			||||||
      andb b₁ (orb b₂ b₃) = orb (andb b₁ b₂) (andb b₁ b₃).
 | 
					 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    solve_bool.
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Definition dist₂ : forall b₁ b₂ b₃,
 | 
					 | 
				
			||||||
      orb b₁ (andb b₂ b₃) = andb (orb b₁ b₂) (orb b₁ b₃).
 | 
					 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    solve_bool.
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Definition max_min : forall b₁ b₂,
 | 
					 | 
				
			||||||
      orb (andb b₁ b₂) b₁ = b₁.
 | 
					 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    solve_bool.
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
End BoolLattice.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
Create HintDb bool_lattice_hints.
 | 
					 | 
				
			||||||
Hint Resolve associativity : bool_lattice_hints.
 | 
					 | 
				
			||||||
(* Hint Resolve (associativity _ _ _)^ : bool_lattice_hints. *)
 | 
					 | 
				
			||||||
Hint Resolve commutativity : bool_lattice_hints.
 | 
					 | 
				
			||||||
Hint Resolve absorption : bool_lattice_hints.
 | 
					 | 
				
			||||||
Hint Resolve idempotency : bool_lattice_hints.
 | 
					 | 
				
			||||||
Hint Resolve left_identity : bool_lattice_hints.
 | 
					 | 
				
			||||||
Hint Resolve right_identity : bool_lattice_hints.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
Hint Resolve
 | 
					 | 
				
			||||||
     associativity
 | 
					 | 
				
			||||||
     and_true and_false
 | 
					 | 
				
			||||||
     dist₁ dist₂ max_min
 | 
					 | 
				
			||||||
  : bool_lattice_hints.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
(** If [B] is a lattice, then [A -> B] is a lattice. *)
 | 
					 | 
				
			||||||
Section fun_lattice.
 | 
					 | 
				
			||||||
  Context {A B : Type}.
 | 
					 | 
				
			||||||
  Context `{BJoin : Join B}.
 | 
					 | 
				
			||||||
  Context `{BMeet : Meet B}.
 | 
					 | 
				
			||||||
  Context `{@Lattice B BJoin BMeet}.
 | 
					 | 
				
			||||||
  Context `{Funext}.
 | 
					 | 
				
			||||||
  Context `{BBottom : Bottom B}.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Global Instance bot_fun : Bottom (A -> B)
 | 
					 | 
				
			||||||
    := fun _ => ⊥.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Global Instance join_fun : Join (A -> B) :=
 | 
					 | 
				
			||||||
    fun (f g : A -> B) (a : A) => (f a) ⊔ (g a).
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Global Instance meet_fun : Meet (A -> B) :=
 | 
					 | 
				
			||||||
    fun (f g : A -> B) (a : A) => (f a) ⊓ (g a).
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Ltac solve_fun :=
 | 
					 | 
				
			||||||
    compute ; intros ; apply path_forall ; intro ;
 | 
					 | 
				
			||||||
    eauto with lattice_hints typeclass_instances.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Create HintDb test1.
 | 
					 | 
				
			||||||
  Lemma associativity_lat `{Lattice A} (x y z : A) :
 | 
					 | 
				
			||||||
    x ⊓ (y ⊓ z) = x ⊓ y ⊓ z.
 | 
					 | 
				
			||||||
  Proof. apply associativity. Defined.
 | 
					 | 
				
			||||||
  Hint Resolve associativity : test1.
 | 
					 | 
				
			||||||
  Hint Resolve associativity_lat : test1.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Global Instance lattice_fun : Lattice (A -> B).
 | 
					 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    repeat (split; try (apply _)).
 | 
					 | 
				
			||||||
    eauto with test1.
 | 
					 | 
				
			||||||
    (* TODO *)
 | 
					 | 
				
			||||||
    all: solve_fun.
 | 
					 | 
				
			||||||
    apply associativity.
 | 
					 | 
				
			||||||
    apply commutativity.
 | 
					 | 
				
			||||||
    apply idempotency. apply _.
 | 
					 | 
				
			||||||
    apply associativity.
 | 
					 | 
				
			||||||
    apply commutativity.
 | 
					 | 
				
			||||||
    apply idempotency. apply _.    
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Global Instance boundedjoinsemilattice_fun
 | 
					 | 
				
			||||||
   `{@BoundedJoinSemiLattice B BJoin BBottom} :
 | 
					 | 
				
			||||||
    BoundedJoinSemiLattice (A -> B).
 | 
					 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    repeat split; try apply _; try solve_fun.
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
End fun_lattice.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
(** If [A] is a lattice and [P] is closed under the lattice operations, then [Σ(x:A), P x] is a lattice. *)
 | 
					 | 
				
			||||||
Section sub_lattice.
 | 
					 | 
				
			||||||
  Context {A : Type} {P : A -> hProp}.
 | 
					 | 
				
			||||||
  Context `{Lattice A}.
 | 
					 | 
				
			||||||
  Context `{Bottom A}.
 | 
					 | 
				
			||||||
  Context {Hmax : forall x y, P x -> P y -> P (x ⊔ y)}.
 | 
					 | 
				
			||||||
  Context {Hmin : forall x y, P x -> P y -> P (x ⊓ y)}.
 | 
					 | 
				
			||||||
  Context {Hbot : P ⊥}.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Definition AP : Type := sig P.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Instance botAP : Bottom AP.
 | 
					 | 
				
			||||||
  Proof. refine (⊥ ↾ _). apply Hbot. Defined.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Instance maxAP : Join AP :=
 | 
					 | 
				
			||||||
    fun x y =>
 | 
					 | 
				
			||||||
      match x, y with
 | 
					 | 
				
			||||||
      | (a ; pa), (b ; pb) => (a ⊔ b ; Hmax a b pa pb)
 | 
					 | 
				
			||||||
      end.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Instance minAP : Meet AP :=
 | 
					 | 
				
			||||||
    fun x y =>
 | 
					 | 
				
			||||||
      match x, y with
 | 
					 | 
				
			||||||
      | (a ; pa), (b ; pb) => (a ⊓ b ; Hmin a b pa pb)
 | 
					 | 
				
			||||||
      end.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Instance hprop_sub : forall c, IsHProp (P c).
 | 
					 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    apply _.
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Ltac solve_sub :=
 | 
					 | 
				
			||||||
    let x := fresh in
 | 
					 | 
				
			||||||
    repeat (intro x ; destruct x)
 | 
					 | 
				
			||||||
    ; simple refine (path_sigma _ _ _ _ _)
 | 
					 | 
				
			||||||
    ; simpl
 | 
					 | 
				
			||||||
    ; try (apply hprop_sub)
 | 
					 | 
				
			||||||
    ; eauto 3 with lattice_hints typeclass_instances.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Global Instance lattice_sub : Lattice AP.
 | 
					 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    repeat (split ; try (apply _ || solve_sub)).
 | 
					 | 
				
			||||||
    apply associativity.
 | 
					 | 
				
			||||||
    apply commutativity.
 | 
					 | 
				
			||||||
    apply idempotency. apply _.
 | 
					 | 
				
			||||||
    apply associativity.
 | 
					 | 
				
			||||||
    apply commutativity.
 | 
					 | 
				
			||||||
    apply idempotency. apply _.
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
End sub_lattice.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
Instance lor : Join hProp := fun X Y => BuildhProp (Trunc (-1) (sum X Y)).
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
Delimit Scope logic_scope with L.
 | 
					 | 
				
			||||||
Notation "A ∨ B" := (lor A B) (at level 20, right associativity) : logic_scope.
 | 
					 | 
				
			||||||
Arguments lor _%L _%L.
 | 
					 | 
				
			||||||
Open Scope logic_scope.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
Instance land : Meet hProp := fun X Y => BuildhProp (prod X Y).
 | 
					 | 
				
			||||||
Instance lfalse : Bottom hProp := False_hp.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
Notation "A ∧ B" := (land A B) (at level 20, right associativity) : logic_scope.
 | 
					 | 
				
			||||||
Arguments land _%L _%L.
 | 
					 | 
				
			||||||
Open Scope logic_scope.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
(** [hProp] is a lattice. *)
 | 
					 | 
				
			||||||
Section hPropLattice.
 | 
					 | 
				
			||||||
  Context `{Univalence}.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Local Ltac lor_intros :=
 | 
					 | 
				
			||||||
    let x := fresh in intro x
 | 
					 | 
				
			||||||
                      ; repeat (strip_truncations ; destruct x as [x | x]).
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Instance lor_commutative : Commutative lor.
 | 
					 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    intros X Y.
 | 
					 | 
				
			||||||
    apply path_iff_hprop ; lor_intros
 | 
					 | 
				
			||||||
    ; apply tr ; auto.
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Instance land_commutative : Commutative land.
 | 
					 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    intros X Y.
 | 
					 | 
				
			||||||
    apply path_hprop.
 | 
					 | 
				
			||||||
    apply equiv_prod_symm.
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Instance lor_associative : Associative lor.
 | 
					 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    intros X Y Z.
 | 
					 | 
				
			||||||
    apply path_iff_hprop ; lor_intros
 | 
					 | 
				
			||||||
    ; apply tr ; auto
 | 
					 | 
				
			||||||
    ; try (left ; apply tr)
 | 
					 | 
				
			||||||
    ; try (right ; apply tr) ; auto.
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Instance land_associative : Associative land.
 | 
					 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    intros X Y Z.
 | 
					 | 
				
			||||||
    symmetry.
 | 
					 | 
				
			||||||
    apply path_hprop.
 | 
					 | 
				
			||||||
    symmetry.
 | 
					 | 
				
			||||||
    apply equiv_prod_assoc.
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Instance lor_idempotent : BinaryIdempotent lor.
 | 
					 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    intros X.
 | 
					 | 
				
			||||||
    apply path_iff_hprop ; lor_intros
 | 
					 | 
				
			||||||
    ; try(refine (tr(inl _))) ; auto.
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Instance land_idempotent : BinaryIdempotent land.
 | 
					 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    intros X.
 | 
					 | 
				
			||||||
    apply path_iff_hprop ; cbn.
 | 
					 | 
				
			||||||
    - intros [a b] ; apply a.
 | 
					 | 
				
			||||||
    - intros a ; apply (pair a a).
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Instance lor_neutrall : LeftIdentity lor lfalse.
 | 
					 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    intros X.
 | 
					 | 
				
			||||||
    apply path_iff_hprop ; lor_intros ; try contradiction
 | 
					 | 
				
			||||||
    ; try (refine (tr(inr _))) ; auto.
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Instance lor_neutralr : RightIdentity lor lfalse.
 | 
					 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    intros X.
 | 
					 | 
				
			||||||
    apply path_iff_hprop ; lor_intros ; try contradiction
 | 
					 | 
				
			||||||
    ; try (refine (tr(inl _))) ; auto.
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Instance absorption_orb_andb : Absorption lor land.
 | 
					 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    intros Z1 Z2.
 | 
					 | 
				
			||||||
    apply path_iff_hprop ; cbn.
 | 
					 | 
				
			||||||
    - intros X ; strip_truncations.
 | 
					 | 
				
			||||||
      destruct X as [Xx | [Xy1 Xy2]] ; assumption.
 | 
					 | 
				
			||||||
    - intros X.
 | 
					 | 
				
			||||||
      apply (tr (inl X)).
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Instance absorption_andb_orb : Absorption land lor.
 | 
					 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    intros Z1 Z2.
 | 
					 | 
				
			||||||
    apply path_iff_hprop ; cbn.
 | 
					 | 
				
			||||||
    - intros [X Y] ; strip_truncations.
 | 
					 | 
				
			||||||
      assumption.
 | 
					 | 
				
			||||||
    - intros X.
 | 
					 | 
				
			||||||
      split.
 | 
					 | 
				
			||||||
      * assumption.
 | 
					 | 
				
			||||||
      * apply (tr (inl X)).
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Global Instance lattice_hprop : Lattice hProp.
 | 
					 | 
				
			||||||
  Proof. repeat (split ; try apply _). Defined.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Global Instance bounded_jsl_hprop : BoundedJoinSemiLattice hProp.
 | 
					 | 
				
			||||||
  Proof. repeat (split ; try apply _). Qed.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Global Instance top_hprop : Top hProp := Unit_hp.
 | 
					 | 
				
			||||||
  Global Instance bounded_msl_hprop : @BoundedSemiLattice hProp (⊓) ⊤.
 | 
					 | 
				
			||||||
  Proof.
 | 
					 | 
				
			||||||
    repeat (split; try apply _); cbv.
 | 
					 | 
				
			||||||
    - intros X. apply path_trunctype ; apply prod_unit_l.
 | 
					 | 
				
			||||||
    - intros X. apply path_trunctype ; apply prod_unit_r.
 | 
					 | 
				
			||||||
  Defined.
 | 
					 | 
				
			||||||
End hPropLattice.
 | 
					 | 
				
			||||||
 
 | 
				
			|||||||
@@ -279,7 +279,7 @@ Section properties.
 | 
				
			|||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Definition union_isIn : forall (A : Type) (a : A) (X Y : T A),
 | 
					  Definition union_isIn : forall (A : Type) (a : A) (X Y : T A),
 | 
				
			||||||
    a ∈ (X ∪ Y) = lor (a ∈ X) (a ∈ Y).
 | 
					    a ∈ (X ∪ Y) = (a ∈ X) ⊔ (a ∈ Y).
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    by (reduce T).
 | 
					    by (reduce T).
 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 
 | 
				
			|||||||
@@ -136,20 +136,14 @@ Section relations.
 | 
				
			|||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    intros A a.
 | 
					    intros A a.
 | 
				
			||||||
    hrecursion.
 | 
					    hrecursion.
 | 
				
			||||||
    - apply False_hp.
 | 
					    - apply ⊥.
 | 
				
			||||||
    - apply (fun a' => merely(a = a')).
 | 
					    - apply (fun a' => merely(a = a')).
 | 
				
			||||||
    - apply lor.
 | 
					    - apply (⊔).
 | 
				
			||||||
      (* TODO *)
 | 
					    - eauto 10 with lattice_hints typeclass_instances.
 | 
				
			||||||
    - eauto with lattice_hints typeclass_instances.
 | 
					    - eauto 10 with lattice_hints typeclass_instances.
 | 
				
			||||||
      apply associativity.
 | 
					    - eauto 10 with lattice_hints typeclass_instances.
 | 
				
			||||||
    - eauto with lattice_hints typeclass_instances.
 | 
					    - eauto 10 with lattice_hints typeclass_instances.
 | 
				
			||||||
      apply commutativity.
 | 
					    - eauto 8 with lattice_hints typeclass_instances.
 | 
				
			||||||
    - eauto with lattice_hints typeclass_instances.
 | 
					 | 
				
			||||||
      apply left_identity.
 | 
					 | 
				
			||||||
    - eauto with lattice_hints typeclass_instances.
 | 
					 | 
				
			||||||
      apply right_identity.
 | 
					 | 
				
			||||||
    - eauto with lattice_hints typeclass_instances.
 | 
					 | 
				
			||||||
      intros. simpl. apply binary_idempotent.
 | 
					 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  (** Subset relation of finite sets. *)
 | 
					  (** Subset relation of finite sets. *)
 | 
				
			||||||
@@ -157,17 +151,13 @@ Section relations.
 | 
				
			|||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    intros A X Y.
 | 
					    intros A X Y.
 | 
				
			||||||
    hrecursion X.
 | 
					    hrecursion X.
 | 
				
			||||||
    - apply Unit_hp.
 | 
					    - apply ⊤.
 | 
				
			||||||
    - apply (fun a => a ∈ Y).
 | 
					    - apply (fun a => a ∈ Y).
 | 
				
			||||||
    - apply land.
 | 
					    - apply (⊓).
 | 
				
			||||||
      (* TODO *)
 | 
					    - eauto 10 with lattice_hints typeclass_instances.
 | 
				
			||||||
    - eauto with lattice_hints typeclass_instances. 
 | 
					    - eauto 10 with lattice_hints typeclass_instances.
 | 
				
			||||||
      apply associativity.
 | 
					    - eauto 10 with lattice_hints typeclass_instances.
 | 
				
			||||||
    - eauto with lattice_hints typeclass_instances.
 | 
					    - eauto 10 with lattice_hints typeclass_instances.
 | 
				
			||||||
      apply commutativity.
 | 
					    - eauto 8 with lattice_hints typeclass_instances.
 | 
				
			||||||
    - apply left_identity.
 | 
					 | 
				
			||||||
    - apply right_identity.
 | 
					 | 
				
			||||||
    - eauto with lattice_hints typeclass_instances.
 | 
					 | 
				
			||||||
      intros. apply binary_idempotent.
 | 
					 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
End relations.
 | 
					End relations.
 | 
				
			||||||
 
 | 
				
			|||||||
@@ -244,7 +244,7 @@ Section properties_membership.
 | 
				
			|||||||
  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_union (ϕ : A -> Bool) : forall X Y : FSet A,
 | 
					  Lemma comprehension_union (ϕ : A -> Bool) : forall X Y : FSet A,
 | 
				
			||||||
      {| (X ∪ Y) | ϕ|} = {|X | ϕ|} ∪ {|Y | ϕ|}.
 | 
					      {| (X ∪ Y) | ϕ|} = {|X | ϕ|} ∪ {|Y | ϕ|}.
 | 
				
			||||||
@@ -261,7 +261,7 @@ Section properties_membership.
 | 
				
			|||||||
      assert (forall c d, ϕ a = c -> ϕ b = d ->
 | 
					      assert (forall c d, ϕ a = c -> ϕ b = d ->
 | 
				
			||||||
                          a ∈ (if ϕ b then {|b|} else ∅)
 | 
					                          a ∈ (if ϕ b then {|b|} else ∅)
 | 
				
			||||||
                          =
 | 
					                          =
 | 
				
			||||||
                          (if ϕ a then BuildhProp (Trunc (-1) (a = b)) else False_hp)) as X.
 | 
					                          (if ϕ a then merely (a = b) else False_hp)) as X.
 | 
				
			||||||
      {
 | 
					      {
 | 
				
			||||||
        intros c d Hc Hd.
 | 
					        intros c d Hc Hd.
 | 
				
			||||||
        destruct c ; destruct d ; rewrite Hc, Hd ; try reflexivity
 | 
					        destruct c ; destruct d ; rewrite Hc, Hd ; try reflexivity
 | 
				
			||||||
@@ -286,7 +286,7 @@ Section properties_membership.
 | 
				
			|||||||
  Context {B : Type}.
 | 
					  Context {B : Type}.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Lemma singleproduct_isIn (a : A) (b : B) (c : A) : forall (Y : FSet B),
 | 
					  Lemma singleproduct_isIn (a : A) (b : B) (c : A) : forall (Y : FSet B),
 | 
				
			||||||
      (a, b) ∈ (single_product c Y) = land (BuildhProp (Trunc (-1) (a = c))) (b ∈ Y).
 | 
					      (a, b) ∈ (single_product c Y) = merely (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.
 | 
				
			||||||
@@ -325,7 +325,7 @@ Section properties_membership.
 | 
				
			|||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Definition product_isIn (a : A) (b : B) (X : FSet A) (Y : FSet B) :
 | 
					  Definition product_isIn (a : A) (b : B) (X : FSet A) (Y : FSet B) :
 | 
				
			||||||
    (a,b) ∈ (product X Y) = land (a ∈ X) (b ∈ Y).
 | 
					    (a,b) ∈ (product X Y) = (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.
 | 
				
			||||||
@@ -664,8 +664,9 @@ Ltac simplify_isIn_d :=
 | 
				
			|||||||
 | 
					
 | 
				
			||||||
Ltac toBool :=
 | 
					Ltac toBool :=
 | 
				
			||||||
  repeat intro;
 | 
					  repeat intro;
 | 
				
			||||||
  apply ext ; intros ;
 | 
					  apply ext; intros;
 | 
				
			||||||
  simplify_isIn_d ; eauto with bool_lattice_hints typeclass_instances.
 | 
					  simplify_isIn_d;
 | 
				
			||||||
 | 
					  eauto 10 with lattice_hints typeclass_instances.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
(** If `A` has decidable equality, then `FSet A` is a lattice *) 
 | 
					(** If `A` has decidable equality, then `FSet A` is a lattice *) 
 | 
				
			||||||
Section set_lattice.
 | 
					Section set_lattice.
 | 
				
			||||||
@@ -680,7 +681,6 @@ Section set_lattice.
 | 
				
			|||||||
    repeat split; try apply _;
 | 
					    repeat split; try apply _;
 | 
				
			||||||
      compute[sg_op meet_is_sg_op meet_fset];
 | 
					      compute[sg_op meet_is_sg_op meet_fset];
 | 
				
			||||||
      toBool.
 | 
					      toBool.
 | 
				
			||||||
    apply associativity.
 | 
					 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
End set_lattice.
 | 
					End set_lattice.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 
 | 
				
			|||||||
@@ -64,7 +64,7 @@ Section quantifiers.
 | 
				
			|||||||
    hinduction.
 | 
					    hinduction.
 | 
				
			||||||
    - apply False_hp.
 | 
					    - apply False_hp.
 | 
				
			||||||
    - apply P.
 | 
					    - apply P.
 | 
				
			||||||
    - apply lor.
 | 
					    - apply (⊔).
 | 
				
			||||||
    - eauto with lattice_hints typeclass_instances.
 | 
					    - eauto with lattice_hints typeclass_instances.
 | 
				
			||||||
    (* TODO eauto with .. *)
 | 
					    (* TODO eauto with .. *)
 | 
				
			||||||
      apply associativity.
 | 
					      apply associativity.
 | 
				
			||||||
 
 | 
				
			|||||||
@@ -189,7 +189,7 @@ Section split.
 | 
				
			|||||||
           (f : {a : A | P a } <~> Fin n + Unit).
 | 
					           (f : {a : A | P a } <~> Fin n + Unit).
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Definition split : exists P' : Sub A, exists b : A,
 | 
					  Definition split : exists P' : Sub A, exists b : A,
 | 
				
			||||||
    prod ({a : A | P' a} <~> Fin n) (forall x, P x = (P' x ∨ merely (x = b))).
 | 
					    prod ({a : A | P' a} <~> Fin n) (forall x, P x = (P' x ⊔ merely (x = b))).
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    pose (fun x : A => sig (fun y : Fin n => x = (f^-1 (inl y)).1)) as P'.
 | 
					    pose (fun x : A => sig (fun y : Fin n => x = (f^-1 (inl y)).1)) as P'.
 | 
				
			||||||
    assert (forall x, IsHProp (P' x)).
 | 
					    assert (forall x, IsHProp (P' x)).
 | 
				
			||||||
 
 | 
				
			|||||||
@@ -7,7 +7,7 @@ Require Import sub prelude FSets list_representation subobjects.k_finite
 | 
				
			|||||||
Fixpoint listExt {A} (ls : list A) : Sub A := fun x =>
 | 
					Fixpoint listExt {A} (ls : list A) : Sub A := fun x =>
 | 
				
			||||||
  match ls with
 | 
					  match ls with
 | 
				
			||||||
  | nil => False_hp
 | 
					  | nil => False_hp
 | 
				
			||||||
  | cons a ls' => BuildhProp (Trunc (-1) (x = a)) ∨ listExt ls' x
 | 
					  | cons a ls' => merely (x = a) ⊔ listExt ls' x
 | 
				
			||||||
  end.
 | 
					  end.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
Fixpoint map {A B} (f : A -> B) (ls : list A) : list B :=
 | 
					Fixpoint map {A B} (f : A -> B) (ls : list A) : list B :=
 | 
				
			||||||
@@ -260,7 +260,7 @@ Section fset_dec_enumerated.
 | 
				
			|||||||
      strip_truncations. apply tr.
 | 
					      strip_truncations. apply tr.
 | 
				
			||||||
      destruct Hls as [ls Hls].
 | 
					      destruct Hls as [ls Hls].
 | 
				
			||||||
      exists (cons a ls). intros b. cbn.
 | 
					      exists (cons a ls). intros b. cbn.
 | 
				
			||||||
      apply (ap (fun z => _ ∨ z) (Hls b)).
 | 
					      apply (ap (fun z => _ ⊔ z) (Hls b)).
 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Definition Kf_enumerated : Kf A -> enumerated A.
 | 
					  Definition Kf_enumerated : Kf A -> enumerated A.
 | 
				
			||||||
 
 | 
				
			|||||||
@@ -10,7 +10,7 @@ Section subobjects.
 | 
				
			|||||||
  Global Instance sub_union : hasUnion Sub := join.
 | 
					  Global Instance sub_union : hasUnion Sub := join.
 | 
				
			||||||
  Global Instance sub_intersection : hasIntersection Sub := meet.
 | 
					  Global Instance sub_intersection : hasIntersection Sub := meet.
 | 
				
			||||||
  Global Instance sub_singleton : hasSingleton Sub A
 | 
					  Global Instance sub_singleton : hasSingleton Sub A
 | 
				
			||||||
    := fun a b => BuildhProp (Trunc (-1) (b = a)).
 | 
					    := fun a b => merely (b = a).
 | 
				
			||||||
  Global Instance sub_membership : hasMembership Sub A := fun a X => X a.
 | 
					  Global Instance sub_membership : hasMembership Sub A := fun a X => X a.
 | 
				
			||||||
  Global Instance sub_comprehension : hasComprehension Sub A
 | 
					  Global Instance sub_comprehension : hasComprehension Sub A
 | 
				
			||||||
    := fun ϕ X a => BuildhProp (X a * (ϕ a = true)).
 | 
					    := fun ϕ X a => BuildhProp (X a * (ϕ a = true)).
 | 
				
			||||||
 
 | 
				
			|||||||
		Reference in New Issue
	
	Block a user