stronger collection
This commit is contained in:
parent
c9f427f737
commit
0578569c09
16
czf.v
16
czf.v
|
@ -311,3 +311,19 @@ Proof.
|
||||||
2: { rewrite Hx. apply (proj2_sig Hϕ'). }
|
2: { rewrite Hx. apply (proj2_sig Hϕ'). }
|
||||||
exists x. cbn. reflexivity.
|
exists x. cbn. reflexivity.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
|
Theorem CZF_collection_stong (a : V) (ϕ : V → V → Prop)
|
||||||
|
`{Proper _ ((≡) ==> (≡) ==> (impl)) ϕ} :
|
||||||
|
(forall (x : V), x ∈ a → { y | ϕ x y }) →
|
||||||
|
exists (b : V), (forall (y : V), y ∈ b → exists x, x ∈ a ∧ ϕ x y ).
|
||||||
|
Proof.
|
||||||
|
intros Hϕ.
|
||||||
|
exists (coll a ϕ Hϕ). intros yy Hyy.
|
||||||
|
destruct a as [A f].
|
||||||
|
destruct Hyy as [y Hy]. simpl in *.
|
||||||
|
exists (f y). split.
|
||||||
|
{ exists y. reflexivity. }
|
||||||
|
rewrite Hy.
|
||||||
|
set (Hϕ':=Hϕ (f y) (ex_intro (λ z, f y ≡ f z) y (reflexivity _))).
|
||||||
|
apply (proj2_sig Hϕ').
|
||||||
|
Qed.
|
||||||
|
|
Loading…
Reference in New Issue