HITs-Examples/FiniteSets/disjunction.v

164 lines
3.6 KiB
Coq
Raw Permalink Normal View History

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-08 15:29:50 +02:00
Require Import lattice notation.
2017-07-31 14:52:41 +02:00
2017-08-03 12:21:34 +02:00
Instance lor : maximum hProp := fun X Y => BuildhProp (Trunc (-1) (sum X Y)).
2017-07-31 14:52:41 +02:00
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-14 16:39:20 +02:00
Local Ltac lor_intros :=
let x := fresh in intro x
; repeat (strip_truncations ; destruct x as [x | x]).
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.
2017-08-14 16:39:20 +02:00
apply path_iff_hprop ; lor_intros
; apply tr ; auto
; try (left ; apply tr)
; try (right ; apply tr) ; auto.
Defined.
2017-07-31 14:52:41 +02:00
2017-08-01 17:12:32 +02:00
Lemma lor_comm : X Y = Y X.
2017-07-31 14:52:41 +02:00
Proof.
2017-08-14 16:39:20 +02:00
apply path_iff_hprop ; lor_intros
; apply tr ; auto.
2017-07-31 14:52:41 +02:00
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.
2017-08-14 16:39:20 +02:00
apply path_iff_hprop ; lor_intros ; try contradiction
; try (refine (tr(inr _))) ; auto.
2017-07-31 14:52:41 +02:00
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.
2017-08-14 16:39:20 +02:00
apply path_iff_hprop ; lor_intros ; try contradiction
; try (refine (tr(inl _))) ; auto.
2017-07-31 14:52:41 +02:00
Defined.
2017-08-01 17:12:32 +02:00
Lemma lor_idem : X X = X.
2017-07-31 14:52:41 +02:00
Proof.
2017-08-14 16:39:20 +02:00
apply path_iff_hprop ; lor_intros
; try(refine (tr(inl _))) ; auto.
2017-07-31 14:54:20 +02:00
Defined.
End lor_props.
2017-08-01 17:12:32 +02:00
2017-08-03 12:21:34 +02:00
Instance land : minimum hProp := fun X Y => BuildhProp (prod X Y).
Instance lfalse : bottom hProp := False_hp.
2017-08-01 17:12:32 +02:00
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}.
2017-08-03 12:21:34 +02:00
Instance lor_commutative : Commutative lor.
2017-08-01 17:12:32 +02:00
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.
2017-08-03 12:21:34 +02:00
Instance lor_neutrall : NeutralL lor lfalse.
2017-08-01 17:12:32 +02:00
Proof.
unfold NeutralL.
intros.
apply lor_nl.
Defined.
2017-08-03 12:21:34 +02:00
Instance lor_neutralr : NeutralR lor lfalse.
2017-08-01 17:12:32 +02:00
Proof.
unfold NeutralR.
intros.
apply lor_nr.
Defined.
2017-08-03 12:21:34 +02:00
Instance absorption_orb_andb : Absorption lor land.
2017-08-01 17:12:32 +02:00
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.
2017-08-03 12:21:34 +02:00
Instance absorption_andb_orb : Absorption land lor.
2017-08-01 17:12:32 +02:00
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-03 12:21:34 +02:00
Global Instance lattice_hprop : Lattice hProp :=
2017-08-01 17:12:32 +02:00
{ commutative_min := _ ;
commutative_max := _ ;
associative_min := _ ;
associative_max := _ ;
idempotent_min := _ ;
idempotent_max := _ ;
2017-08-03 12:21:34 +02:00
neutralL_max := _ ;
neutralR_max := _ ;
2017-08-01 17:12:32 +02:00
absorption_min_max := _ ;
absorption_max_min := _
}.
2017-08-02 15:45:12 +02:00
End hPropLattice.