extensionality
This commit is contained in:
14
czf.v
14
czf.v
@@ -184,6 +184,12 @@ Proof.
|
|||||||
rewrite <- Ha. rewrite <- Hb. apply Hs.
|
rewrite <- Ha. rewrite <- Hb. apply Hs.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
|
Instance subseteq_reflexive : Reflexive (subseteq : relation V).
|
||||||
|
Proof. intros ?. unfold subseteq, subseteq_V. auto. Qed.
|
||||||
|
|
||||||
|
Instance subseteq_transitive : Transitive (subseteq : relation V).
|
||||||
|
Proof. intros ???. unfold subseteq, subseteq_V. auto. Qed.
|
||||||
|
|
||||||
Lemma V_ext (a b : V) :
|
Lemma V_ext (a b : V) :
|
||||||
a ⊆ b → b ⊆ a → a ≡ b.
|
a ⊆ b → b ⊆ a → a ≡ b.
|
||||||
Proof.
|
Proof.
|
||||||
@@ -200,7 +206,13 @@ Qed.
|
|||||||
(** * Axioms *)
|
(** * Axioms *)
|
||||||
|
|
||||||
(** Extensionality *)
|
(** Extensionality *)
|
||||||
(* ... *)
|
Theorem CZF_extensionality (a b : V) :
|
||||||
|
a ⊆ b ∧ b ⊆ a ↔ a ≡ b.
|
||||||
|
Proof.
|
||||||
|
split.
|
||||||
|
- intros [? ?]. apply V_ext; auto.
|
||||||
|
- intros ->. split; auto.
|
||||||
|
Qed.
|
||||||
|
|
||||||
(** Paring *)
|
(** Paring *)
|
||||||
Lemma CZF_pairing (x y : V) :
|
Lemma CZF_pairing (x y : V) :
|
||||||
|
|||||||
Reference in New Issue
Block a user