Library UniMath.MoreFoundations.Propositions
Require Export UniMath.MoreFoundations.Notations.
Require Export UniMath.MoreFoundations.Tactics.
Require Export UniMath.MoreFoundations.DecidablePropositions.
Local Open Scope logic.
Lemma ishinh_irrel {X:UU} (x:X) (x':∥X∥) : hinhpr x = x'.
Show proof.
Corollary squash_to_hProp {X:UU} {Q:hProp} : ∥ X ∥ -> (X -> Q) -> Q.
Show proof.
Lemma hdisj_impl_1 {P Q : hProp} : P∨Q -> (Q->P) -> P.
Show proof.
Lemma hdisj_impl_2 {P Q : hProp} : P∨Q -> (P->Q) -> Q.
Show proof.
Definition weqlogeq (P Q : hProp) : (P = Q) ≃ (P ⇔ Q).
Show proof.
Lemma decidable_proof_by_contradiction {P:hProp} : decidable P -> ¬ ¬ P -> P.
Show proof.
Lemma proof_by_contradiction {P:hProp} : LEM -> ¬ ¬ P -> P.
Show proof.
Lemma dneg_elim_to_LEM : (∏ P:hProp, ¬ ¬ P -> P) -> LEM.
Show proof.
Lemma negforall_to_existsneg {X:UU} (P:X->hProp) : LEM -> (¬ ∀ x, P x) -> (∃ x, ¬ (P x)).
Show proof.
Lemma negimpl_to_conj (P Q:hProp) : LEM -> ( ¬ (P ⇒ Q) -> P ∧ ¬ Q ).
Show proof.
Definition hrel_set (X : hSet) : hSet := hSetpair (hrel X) (isaset_hrel X).
Lemma isaprop_assume_it_is {X : UU} : (X -> isaprop X) -> isaprop X.
Show proof.
Lemma isaproptotal2 {X : UU} (P : X → UU) :
isPredicate P →
(∏ x y : X, P x → P y → x = y) →
isaprop (∑ x : X, P x).
Show proof.
Lemma squash_rec {X : UU} (P : ∥ X ∥ -> hProp) : (∏ x, P (hinhpr x)) -> (∏ x, P x).
Show proof.
Lemma logeq_if_both_true (P Q : hProp) : P -> Q -> ( P ⇔ Q ).
Show proof.
Lemma logeq_if_both_false (P Q : hProp) : ¬P -> ¬Q -> ( P ⇔ Q ).
Show proof.
Definition proofirrelevance_hProp (X : hProp) : isProofIrrelevant X
:= proofirrelevance X (propproperty X).
Ltac induction_hProp x y := induction (proofirrelevance_hProp _ x y).
Definition iscontr_hProp (X:UU) : hProp := hProppair (iscontr X) (isapropiscontr X).
Notation "'∃!' x .. y , P"
:= (iscontr_hProp (∑ x, .. (∑ y, P) ..))
(at level 200, x binder, y binder, right associativity) : type_scope.
Require Export UniMath.MoreFoundations.Tactics.
Require Export UniMath.MoreFoundations.DecidablePropositions.
Local Open Scope logic.
Lemma ishinh_irrel {X:UU} (x:X) (x':∥X∥) : hinhpr x = x'.
Show proof.
Corollary squash_to_hProp {X:UU} {Q:hProp} : ∥ X ∥ -> (X -> Q) -> Q.
Show proof.
Lemma hdisj_impl_1 {P Q : hProp} : P∨Q -> (Q->P) -> P.
Show proof.
Lemma hdisj_impl_2 {P Q : hProp} : P∨Q -> (P->Q) -> Q.
Show proof.
Definition weqlogeq (P Q : hProp) : (P = Q) ≃ (P ⇔ Q).
Show proof.
intros.
apply weqimplimpl.
- intro e. induction e. apply isrefl_logeq.
- intro c. apply hPropUnivalence.
+ exact (pr1 c).
+ exact (pr2 c).
- apply isasethProp.
- apply propproperty.
apply weqimplimpl.
- intro e. induction e. apply isrefl_logeq.
- intro c. apply hPropUnivalence.
+ exact (pr1 c).
+ exact (pr2 c).
- apply isasethProp.
- apply propproperty.
Lemma decidable_proof_by_contradiction {P:hProp} : decidable P -> ¬ ¬ P -> P.
Show proof.
Lemma proof_by_contradiction {P:hProp} : LEM -> ¬ ¬ P -> P.
Show proof.
Lemma dneg_elim_to_LEM : (∏ P:hProp, ¬ ¬ P -> P) -> LEM.
Show proof.
intros dne. intros P. simple refine (dne (_,,_) _).
simpl. intros n.
assert (q : ¬ (P ∨ ¬ P)).
{ now apply weqnegtonegishinh. }
assert (r := fromnegcoprod_prop q).
exact (pr2 r (pr1 r)).
simpl. intros n.
assert (q : ¬ (P ∨ ¬ P)).
{ now apply weqnegtonegishinh. }
assert (r := fromnegcoprod_prop q).
exact (pr2 r (pr1 r)).
Lemma negforall_to_existsneg {X:UU} (P:X->hProp) : LEM -> (¬ ∀ x, P x) -> (∃ x, ¬ (P x)).
Show proof.
intros lem nf. apply (proof_by_contradiction lem); intro c. use nf; clear nf. intro x.
assert (q := neghexisttoforallneg _ c x); clear c; simpl in q.
exact (proof_by_contradiction lem q).
assert (q := neghexisttoforallneg _ c x); clear c; simpl in q.
exact (proof_by_contradiction lem q).
Lemma negimpl_to_conj (P Q:hProp) : LEM -> ( ¬ (P ⇒ Q) -> P ∧ ¬ Q ).
Show proof.
intros lem ni. assert (r := negforall_to_existsneg _ lem ni); clear lem ni.
apply (squash_to_hProp r); clear r; intros [p nq]. exact (p,,nq).
apply (squash_to_hProp r); clear r; intros [p nq]. exact (p,,nq).
Definition hrel_set (X : hSet) : hSet := hSetpair (hrel X) (isaset_hrel X).
Lemma isaprop_assume_it_is {X : UU} : (X -> isaprop X) -> isaprop X.
Show proof.
Lemma isaproptotal2 {X : UU} (P : X → UU) :
isPredicate P →
(∏ x y : X, P x → P y → x = y) →
isaprop (∑ x : X, P x).
Show proof.
intros HP Heq.
apply invproofirrelevance.
intros [x p] [y q].
induction (Heq x y p q).
induction (iscontrpr1 (HP x p q)).
apply idpath.
apply invproofirrelevance.
intros [x p] [y q].
induction (Heq x y p q).
induction (iscontrpr1 (HP x p q)).
apply idpath.
Lemma squash_rec {X : UU} (P : ∥ X ∥ -> hProp) : (∏ x, P (hinhpr x)) -> (∏ x, P x).
Show proof.
intros xp x'. simple refine (hinhuniv _ x'). intro x.
assert (q : hinhpr x = x').
{ apply propproperty. }
induction q. exact (xp x).
assert (q : hinhpr x = x').
{ apply propproperty. }
induction q. exact (xp x).
Lemma logeq_if_both_true (P Q : hProp) : P -> Q -> ( P ⇔ Q ).
Show proof.
intros p q.
split.
- intros _. exact q.
- intros _. exact p.
split.
- intros _. exact q.
- intros _. exact p.
Lemma logeq_if_both_false (P Q : hProp) : ¬P -> ¬Q -> ( P ⇔ Q ).
Show proof.
intros np nq.
split.
- intros p. apply fromempty. exact (np p).
- intros q. apply fromempty. exact (nq q).
split.
- intros p. apply fromempty. exact (np p).
- intros q. apply fromempty. exact (nq q).
Definition proofirrelevance_hProp (X : hProp) : isProofIrrelevant X
:= proofirrelevance X (propproperty X).
Ltac induction_hProp x y := induction (proofirrelevance_hProp _ x y).
Definition iscontr_hProp (X:UU) : hProp := hProppair (iscontr X) (isapropiscontr X).
Notation "'∃!' x .. y , P"
:= (iscontr_hProp (∑ x, .. (∑ y, P) ..))
(at level 200, x binder, y binder, right associativity) : type_scope.
Various algebraic properties of hProp
We first state the algebraic properties as bi-implications
Lemma islogeqassochconj {P Q R : hProp} : (P ∧ Q) ∧ R <-> P ∧ (Q ∧ R).
Show proof.
split.
- intros PQR.
exact (pr1 (pr1 PQR),,(pr2 (pr1 PQR),,pr2 PQR)).
- intros PQR.
exact ((pr1 PQR,,pr1 (pr2 PQR)),,pr2 (pr2 PQR)).
- intros PQR.
exact (pr1 (pr1 PQR),,(pr2 (pr1 PQR),,pr2 PQR)).
- intros PQR.
exact ((pr1 PQR,,pr1 (pr2 PQR)),,pr2 (pr2 PQR)).
Lemma islogeqcommhconj {P Q : hProp} : P ∧ Q <-> Q ∧ P.
Show proof.
Lemma islogeqassochdisj {P Q R : hProp} : (P ∨ Q) ∨ R <-> P ∨ (Q ∨ R).
Show proof.
split.
- apply hinhuniv; intros hPQR.
induction hPQR as [hPQ|hR].
+ use (hinhuniv _ hPQ); clear hPQ; intros hPQ.
induction hPQ as [hP|hQ].
* exact (hinhpr (ii1 hP)).
* exact (hinhpr (ii2 (hinhpr (ii1 hQ)))).
+ exact (hinhpr (ii2 (hinhpr (ii2 hR)))).
- apply hinhuniv; intros hPQR.
induction hPQR as [hP|hQR].
+ exact (hinhpr (ii1 (hinhpr (ii1 hP)))).
+ use (hinhuniv _ hQR); clear hQR; intros hQR.
induction hQR as [hQ|hR].
* exact (hinhpr (ii1 (hinhpr (ii2 hQ)))).
* exact (hinhpr (ii2 hR)).
- apply hinhuniv; intros hPQR.
induction hPQR as [hPQ|hR].
+ use (hinhuniv _ hPQ); clear hPQ; intros hPQ.
induction hPQ as [hP|hQ].
* exact (hinhpr (ii1 hP)).
* exact (hinhpr (ii2 (hinhpr (ii1 hQ)))).
+ exact (hinhpr (ii2 (hinhpr (ii2 hR)))).
- apply hinhuniv; intros hPQR.
induction hPQR as [hP|hQR].
+ exact (hinhpr (ii1 (hinhpr (ii1 hP)))).
+ use (hinhuniv _ hQR); clear hQR; intros hQR.
induction hQR as [hQ|hR].
* exact (hinhpr (ii1 (hinhpr (ii2 hQ)))).
* exact (hinhpr (ii2 hR)).
Lemma islogeqhconj_absorb_hdisj {P Q : hProp} : P ∧ (P ∨ Q) <-> P.
Show proof.
Lemma islogeqhdisj_absorb_hconj {P Q : hProp} : P ∨ (P ∧ Q) <-> P.
Show proof.
split.
- apply hinhuniv; intros hPPQ.
induction hPPQ as [hP|hPQ].
+ exact hP.
+ exact (pr1 hPQ).
- intros hP; apply (hinhpr (ii1 hP)).
- apply hinhuniv; intros hPPQ.
induction hPPQ as [hP|hPQ].
+ exact hP.
+ exact (pr1 hPQ).
- intros hP; apply (hinhpr (ii1 hP)).
Lemma islogeqhfalse_hdisj {P : hProp} : ∅ ∨ P <-> P.
Show proof.
split.
- apply hinhuniv; intros hPPQ.
induction hPPQ as [hF|hP].
+ induction hF.
+ exact hP.
- intros hP; apply (hinhpr (ii2 hP)).
- apply hinhuniv; intros hPPQ.
induction hPPQ as [hF|hP].
+ induction hF.
+ exact hP.
- intros hP; apply (hinhpr (ii2 hP)).
Lemma islogeqhhtrue_hconj {P : hProp} : htrue ∧ P <-> P.
Show proof.
We now turn these into equalities using univalence for propositions
Lemma isassoc_hconj (P Q R : hProp) : ((P ∧ Q) ∧ R) = (P ∧ (Q ∧ R)).
Show proof.
Lemma iscomm_hconj (P Q : hProp) : (P ∧ Q) = (Q ∧ P).
Show proof.
Lemma isassoc_hdisj (P Q R : hProp) : ((P ∨ Q) ∨ R) = (P ∨ (Q ∨ R)).
Show proof.
Lemma iscomm_hdisj (P Q : hProp) : (P ∨ Q) = (Q ∨ P).
Show proof.
Lemma hconj_absorb_hdisj (P Q : hProp) : (P ∧ (P ∨ Q)) = P.
Show proof.
Lemma hdisj_absorb_hconj (P Q : hProp) : (P ∨ (P ∧ Q)) = P.
Show proof.
Lemma hfalse_hdisj (P : hProp) : (∅ ∨ P) = P.
Show proof.
Lemma htrue_hconj (P : hProp) : (htrue ∧ P) = P.
Show proof.
End hProp_logic.
Show proof.
Lemma iscomm_hconj (P Q : hProp) : (P ∧ Q) = (Q ∧ P).
Show proof.
Lemma isassoc_hdisj (P Q R : hProp) : ((P ∨ Q) ∨ R) = (P ∨ (Q ∨ R)).
Show proof.
Lemma iscomm_hdisj (P Q : hProp) : (P ∨ Q) = (Q ∨ P).
Show proof.
Lemma hconj_absorb_hdisj (P Q : hProp) : (P ∧ (P ∨ Q)) = P.
Show proof.
Lemma hdisj_absorb_hconj (P Q : hProp) : (P ∨ (P ∧ Q)) = P.
Show proof.
Lemma hfalse_hdisj (P : hProp) : (∅ ∨ P) = P.
Show proof.
Lemma htrue_hconj (P : hProp) : (htrue ∧ P) = P.
Show proof.
End hProp_logic.