HITs-Examples/FiniteSets/disjunction.v

76 lines
1.9 KiB
Coq
Raw 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.
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 15:12:59 +02:00
Section lor_props.
2017-07-31 14:52:41 +02:00
Variable X Y Z : hProp.
Context `{Univalence}.
2017-08-01 15:12:59 +02:00
Theorem 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 [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)).
* 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)))).
Defined.
2017-08-01 15:12:59 +02:00
Theorem 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 15:12:59 +02:00
Theorem 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 15:12:59 +02:00
Theorem 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 15:12:59 +02:00
Theorem 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.