2017-08-01 15:12:59 +02:00
|
|
|
|
(* Logical disjunction in HoTT (see ch. 3 of the book) *)
|
2017-07-31 14:52:41 +02:00
|
|
|
|
Require Import HoTT.
|
2017-08-01 17:12:32 +02:00
|
|
|
|
Require Import lattice.
|
2017-07-31 14:52:41 +02:00
|
|
|
|
|
|
|
|
|
Definition lor (X Y : hProp) : hProp := BuildhProp (Trunc (-1) (sum X Y)).
|
|
|
|
|
|
2017-08-01 15:12:59 +02:00
|
|
|
|
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.
|
2017-07-31 14:52:41 +02:00
|
|
|
|
|
2017-08-01 17:25:57 +02:00
|
|
|
|
Section lor_props.
|
2017-07-31 14:52:41 +02:00
|
|
|
|
Context `{Univalence}.
|
2017-08-01 17:25:57 +02:00
|
|
|
|
Variable X Y Z : hProp.
|
2017-08-01 15:12:59 +02:00
|
|
|
|
|
2017-08-01 17:12:32 +02:00
|
|
|
|
Lemma lor_assoc : (X ∨ Y) ∨ Z = X ∨ Y ∨ Z.
|
2017-07-31 14:52:41 +02:00
|
|
|
|
Proof.
|
|
|
|
|
apply path_iff_hprop ; cbn.
|
|
|
|
|
* simple refine (Trunc_ind _ _).
|
|
|
|
|
intros [xy | z] ; cbn.
|
|
|
|
|
+ simple refine (Trunc_ind _ _ xy).
|
|
|
|
|
intros [x | y].
|
|
|
|
|
++ apply (tr (inl x)).
|
|
|
|
|
++ apply (tr (inr (tr (inl y)))).
|
|
|
|
|
+ apply (tr (inr (tr (inr z)))).
|
2017-08-01 17:12:32 +02:00
|
|
|
|
* simple refine (Trunc_ind _ _).
|
|
|
|
|
intros [x | yz] ; cbn.
|
|
|
|
|
+ apply (tr (inl (tr (inl x)))).
|
|
|
|
|
+ simple refine (Trunc_ind _ _ yz).
|
|
|
|
|
intros [y | z].
|
|
|
|
|
++ apply (tr (inl (tr (inr y)))).
|
|
|
|
|
++ apply (tr (inr z)).
|
2017-07-31 14:52:41 +02:00
|
|
|
|
Defined.
|
|
|
|
|
|
2017-08-01 17:12:32 +02:00
|
|
|
|
Lemma lor_comm : X ∨ Y = Y ∨ X.
|
2017-07-31 14:52:41 +02:00
|
|
|
|
Proof.
|
|
|
|
|
apply path_iff_hprop ; cbn.
|
|
|
|
|
* simple refine (Trunc_ind _ _).
|
|
|
|
|
intros [x | y].
|
|
|
|
|
+ apply (tr (inr x)).
|
|
|
|
|
+ apply (tr (inl y)).
|
|
|
|
|
* simple refine (Trunc_ind _ _).
|
|
|
|
|
intros [y | x].
|
|
|
|
|
+ apply (tr (inr y)).
|
|
|
|
|
+ apply (tr (inl x)).
|
|
|
|
|
Defined.
|
|
|
|
|
|
2017-08-01 17:25:57 +02:00
|
|
|
|
Lemma lor_nl : False_hp ∨ X = X.
|
2017-07-31 14:52:41 +02:00
|
|
|
|
Proof.
|
|
|
|
|
apply path_iff_hprop ; cbn.
|
|
|
|
|
* simple refine (Trunc_ind _ _).
|
|
|
|
|
intros [ | x].
|
|
|
|
|
+ apply Empty_rec.
|
|
|
|
|
+ apply x.
|
|
|
|
|
* apply (fun x => tr (inr x)).
|
|
|
|
|
Defined.
|
|
|
|
|
|
2017-08-01 17:25:57 +02:00
|
|
|
|
Lemma lor_nr : X ∨ False_hp = X.
|
2017-07-31 14:52:41 +02:00
|
|
|
|
Proof.
|
|
|
|
|
apply path_iff_hprop ; cbn.
|
|
|
|
|
* simple refine (Trunc_ind _ _).
|
|
|
|
|
intros [x | ].
|
|
|
|
|
+ apply x.
|
|
|
|
|
+ apply Empty_rec.
|
|
|
|
|
* apply (fun x => tr (inl x)).
|
|
|
|
|
Defined.
|
|
|
|
|
|
2017-08-01 17:12:32 +02:00
|
|
|
|
Lemma lor_idem : X ∨ X = X.
|
2017-07-31 14:52:41 +02:00
|
|
|
|
Proof.
|
|
|
|
|
apply path_iff_hprop ; cbn.
|
|
|
|
|
- simple refine (Trunc_ind _ _).
|
|
|
|
|
intros [x | x] ; apply x.
|
|
|
|
|
- apply (fun x => tr (inl x)).
|
2017-07-31 14:54:20 +02:00
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
|
|
End lor_props.
|
2017-08-01 17:12:32 +02:00
|
|
|
|
|
|
|
|
|
Definition land (X Y : hProp) : hProp := BuildhProp (prod X Y).
|
|
|
|
|
|
|
|
|
|
Notation "A ∧ B" := (land A B) (at level 20, right associativity) : logic_scope.
|
|
|
|
|
Arguments land _%L _%L.
|
|
|
|
|
Open Scope logic_scope.
|
|
|
|
|
|
|
|
|
|
Section hPropLattice.
|
|
|
|
|
Context `{Univalence}.
|
|
|
|
|
|
|
|
|
|
Instance lor_commutative : Commutative lor.
|
|
|
|
|
Proof.
|
|
|
|
|
unfold Commutative.
|
|
|
|
|
intros.
|
|
|
|
|
apply lor_comm.
|
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
|
|
Instance land_commutative : Commutative land.
|
|
|
|
|
Proof.
|
|
|
|
|
unfold Commutative, land.
|
|
|
|
|
intros.
|
|
|
|
|
apply path_hprop.
|
|
|
|
|
apply equiv_prod_symm.
|
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
|
|
Instance lor_associative : Associative lor.
|
|
|
|
|
Proof.
|
|
|
|
|
unfold Associative.
|
|
|
|
|
intros.
|
|
|
|
|
apply lor_assoc.
|
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
|
|
Instance land_associative : Associative land.
|
|
|
|
|
Proof.
|
|
|
|
|
unfold Associative, land.
|
|
|
|
|
intros.
|
|
|
|
|
symmetry.
|
|
|
|
|
apply path_hprop.
|
|
|
|
|
apply equiv_prod_assoc.
|
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
|
|
Instance lor_idempotent : Idempotent lor.
|
|
|
|
|
Proof.
|
|
|
|
|
unfold Idempotent.
|
|
|
|
|
intros.
|
|
|
|
|
apply lor_idem.
|
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
|
|
Instance land_idempotent : Idempotent land.
|
|
|
|
|
Proof.
|
|
|
|
|
unfold Idempotent, land.
|
|
|
|
|
intros.
|
|
|
|
|
apply path_iff_hprop ; cbn.
|
|
|
|
|
- intros [a b] ; apply a.
|
|
|
|
|
- intros a ; apply (pair a a).
|
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
|
|
Instance lor_neutrall : NeutralL lor False_hp.
|
|
|
|
|
Proof.
|
|
|
|
|
unfold NeutralL.
|
|
|
|
|
intros.
|
|
|
|
|
apply lor_nl.
|
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
|
|
Instance lor_neutralr : NeutralR lor False_hp.
|
|
|
|
|
Proof.
|
|
|
|
|
unfold NeutralR.
|
|
|
|
|
intros.
|
|
|
|
|
apply lor_nr.
|
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
|
|
Instance bool_absorption_orb_andb : Absorption lor land.
|
|
|
|
|
Proof.
|
|
|
|
|
unfold Absorption.
|
|
|
|
|
intros.
|
|
|
|
|
apply path_iff_hprop ; cbn.
|
|
|
|
|
- intros X ; strip_truncations.
|
|
|
|
|
destruct X as [Xx | [Xy1 Xy2]] ; assumption.
|
|
|
|
|
- intros X.
|
|
|
|
|
apply (tr (inl X)).
|
|
|
|
|
Defined.
|
|
|
|
|
|
|
|
|
|
Instance bool_absorption_andb_orb : Absorption land lor.
|
|
|
|
|
Proof.
|
|
|
|
|
unfold Absorption.
|
|
|
|
|
intros.
|
|
|
|
|
apply path_iff_hprop ; cbn.
|
|
|
|
|
- intros [X Y] ; strip_truncations.
|
|
|
|
|
assumption.
|
|
|
|
|
- intros X.
|
|
|
|
|
split.
|
|
|
|
|
* assumption.
|
|
|
|
|
* apply (tr (inl X)).
|
|
|
|
|
Defined.
|
|
|
|
|
|
2017-08-01 17:25:57 +02:00
|
|
|
|
Global Instance lattice_hprop : Lattice land lor False_hp :=
|
2017-08-01 17:12:32 +02:00
|
|
|
|
{ commutative_min := _ ;
|
|
|
|
|
commutative_max := _ ;
|
|
|
|
|
associative_min := _ ;
|
|
|
|
|
associative_max := _ ;
|
|
|
|
|
idempotent_min := _ ;
|
|
|
|
|
idempotent_max := _ ;
|
|
|
|
|
neutralL_min := _ ;
|
|
|
|
|
neutralR_min := _ ;
|
|
|
|
|
absorption_min_max := _ ;
|
|
|
|
|
absorption_max_min := _
|
|
|
|
|
}.
|
|
|
|
|
|
2017-08-01 17:25:57 +02:00
|
|
|
|
End hPropLattice.
|
|
|
|
|
|
|
|
|
|
Hint Resolve
|
|
|
|
|
commutative_min commutative_max associative_min associative_max
|
|
|
|
|
idempotent_min idempotent_max
|
|
|
|
|
neutralL_min neutralR_min
|
|
|
|
|
absorption_min_max absorption_max_min
|
|
|
|
|
: lattice_hints.
|