HomeHome Intuitionistic Logic Explorer
Theorem List (p. 108 of 108)
< Previous  Wrap >
Bad symbols? Try the
GIF version.

Mirrors  >  Metamath Home Page  >  ILE Home Page  >  Theorem List Contents  >  Recent Proofs       This page: Page List

Theorem List for Intuitionistic Logic Explorer - 10701-10795   *Has distinct variable group(s)
TypeLabelDescription
Statement
 
Theorembj-zfpair2 10701 Proof of zfpair2 3965 using only bounded separation. (Contributed by BJ, 5-Oct-2019.) (Proof modification is discouraged.)
{𝑥, 𝑦} ∈ V
 
Theorembj-prexg 10702 Proof of prexg 3966 using only bounded separation. (Contributed by BJ, 5-Oct-2019.) (Proof modification is discouraged.)
((𝐴𝑉𝐵𝑊) → {𝐴, 𝐵} ∈ V)
 
Theorembj-snexg 10703 snexg 3956 from bounded separation. (Contributed by BJ, 5-Oct-2019.) (Proof modification is discouraged.)
(𝐴𝑉 → {𝐴} ∈ V)
 
Theorembj-snex 10704 snex 3957 from bounded separation. (Contributed by BJ, 5-Oct-2019.) (Proof modification is discouraged.)
𝐴 ∈ V       {𝐴} ∈ V
 
Theorembj-sels 10705* If a class is a set, then it is a member of a set. (Copied from set.mm.) (Contributed by BJ, 3-Apr-2019.)
(𝐴𝑉 → ∃𝑥 𝐴𝑥)
 
Theorembj-axun2 10706* axun2 4190 from bounded separation. (Contributed by BJ, 15-Oct-2019.) (Proof modification is discouraged.)
𝑦𝑧(𝑧𝑦 ↔ ∃𝑤(𝑧𝑤𝑤𝑥))
 
Theorembj-uniex2 10707* uniex2 4191 from bounded separation. (Contributed by BJ, 15-Oct-2019.) (Proof modification is discouraged.)
𝑦 𝑦 = 𝑥
 
Theorembj-uniex 10708 uniex 4192 from bounded separation. (Contributed by BJ, 13-Nov-2019.) (Proof modification is discouraged.)
𝐴 ∈ V        𝐴 ∈ V
 
Theorembj-uniexg 10709 uniexg 4193 from bounded separation. (Contributed by BJ, 13-Nov-2019.) (Proof modification is discouraged.)
(𝐴𝑉 𝐴 ∈ V)
 
Theorembj-unex 10710 unex 4194 from bounded separation. (Contributed by BJ, 13-Nov-2019.) (Proof modification is discouraged.)
𝐴 ∈ V    &   𝐵 ∈ V       (𝐴𝐵) ∈ V
 
Theorembdunexb 10711 Bounded version of unexb 4195. (Contributed by BJ, 13-Nov-2019.) (Proof modification is discouraged.)
BOUNDED 𝐴    &   BOUNDED 𝐵       ((𝐴 ∈ V ∧ 𝐵 ∈ V) ↔ (𝐴𝐵) ∈ V)
 
Theorembj-unexg 10712 unexg 4196 from bounded separation. (Contributed by BJ, 13-Nov-2019.) (Proof modification is discouraged.)
((𝐴𝑉𝐵𝑊) → (𝐴𝐵) ∈ V)
 
Theorembj-sucexg 10713 sucexg 4242 from bounded separation. (Contributed by BJ, 13-Nov-2019.) (Proof modification is discouraged.)
(𝐴𝑉 → suc 𝐴 ∈ V)
 
Theorembj-sucex 10714 sucex 4243 from bounded separation. (Contributed by BJ, 13-Nov-2019.) (Proof modification is discouraged.)
𝐴 ∈ V       suc 𝐴 ∈ V
 
6.2.5.1  Delta_0-classical logic
 
Axiomax-bj-d0cl 10715 Axiom for Δ0-classical logic. (Contributed by BJ, 2-Jan-2020.)
BOUNDED 𝜑       DECID 𝜑
 
Theorembj-notbi 10716 Equivalence property for negation. TODO: minimize all theorems using notbid 624 and notbii 626. (Contributed by BJ, 27-Jan-2020.) (Proof modification is discouraged.)
((𝜑𝜓) → (¬ 𝜑 ↔ ¬ 𝜓))
 
Theorembj-notbii 10717 Inference associated with bj-notbi 10716. (Contributed by BJ, 27-Jan-2020.) (Proof modification is discouraged.)
(𝜑𝜓)       𝜑 ↔ ¬ 𝜓)
 
Theorembj-notbid 10718 Deduction form of bj-notbi 10716. (Contributed by BJ, 27-Jan-2020.) (Proof modification is discouraged.)
(𝜑 → (𝜓𝜒))       (𝜑 → (¬ 𝜓 ↔ ¬ 𝜒))
 
Theorembj-dcbi 10719 Equivalence property for DECID. TODO: solve conflict with dcbi 877; minimize dcbii 780 and dcbid 781 with it, as well as theorems using those. (Contributed by BJ, 27-Jan-2020.) (Proof modification is discouraged.)
((𝜑𝜓) → (DECID 𝜑DECID 𝜓))
 
Theorembj-d0clsepcl 10720 Δ0-classical logic and separation implies classical logic. (Contributed by BJ, 2-Jan-2020.) (Proof modification is discouraged.)
DECID 𝜑
 
6.2.5.2  Inductive classes and the class of natural numbers (finite ordinals)
 
Syntaxwind 10721 Syntax for inductive classes.
wff Ind 𝐴
 
Definitiondf-bj-ind 10722* Define the property of being an inductive class. (Contributed by BJ, 30-Nov-2019.)
(Ind 𝐴 ↔ (∅ ∈ 𝐴 ∧ ∀𝑥𝐴 suc 𝑥𝐴))
 
Theorembj-indsuc 10723 A direct consequence of the definition of Ind. (Contributed by BJ, 30-Nov-2019.)
(Ind 𝐴 → (𝐵𝐴 → suc 𝐵𝐴))
 
Theorembj-indeq 10724 Equality property for Ind. (Contributed by BJ, 30-Nov-2019.)
(𝐴 = 𝐵 → (Ind 𝐴 ↔ Ind 𝐵))
 
Theorembj-bdind 10725 Boundedness of the formula "the setvar 𝑥 is an inductive class". (Contributed by BJ, 30-Nov-2019.)
BOUNDED Ind 𝑥
 
Theorembj-indint 10726* The property of being an inductive class is closed under intersections. (Contributed by BJ, 30-Nov-2019.)
Ind {𝑥𝐴 ∣ Ind 𝑥}
 
Theorembj-indind 10727* If 𝐴 is inductive and 𝐵 is "inductive in 𝐴", then (𝐴𝐵) is inductive. (Contributed by BJ, 25-Oct-2020.)
((Ind 𝐴 ∧ (∅ ∈ 𝐵 ∧ ∀𝑥𝐴 (𝑥𝐵 → suc 𝑥𝐵))) → Ind (𝐴𝐵))
 
Theorembj-dfom 10728 Alternate definition of ω, as the intersection of all the inductive sets. Proposal: make this the definition. (Contributed by BJ, 30-Nov-2019.)
ω = {𝑥 ∣ Ind 𝑥}
 
Theorembj-omind 10729 ω is an inductive class. (Contributed by BJ, 30-Nov-2019.)
Ind ω
 
Theorembj-omssind 10730 ω is included in all the inductive sets (but for the moment, we cannot prove that it is included in all the inductive classes). (Contributed by BJ, 30-Nov-2019.) (Proof modification is discouraged.)
(𝐴𝑉 → (Ind 𝐴 → ω ⊆ 𝐴))
 
Theorembj-ssom 10731* A characterization of subclasses of ω. (Contributed by BJ, 30-Nov-2019.) (Proof modification is discouraged.)
(∀𝑥(Ind 𝑥𝐴𝑥) ↔ 𝐴 ⊆ ω)
 
Theorembj-om 10732* A set is equal to ω if and only if it is the smallest inductive set. (Contributed by BJ, 30-Nov-2019.) (Proof modification is discouraged.)
(𝐴𝑉 → (𝐴 = ω ↔ (Ind 𝐴 ∧ ∀𝑥(Ind 𝑥𝐴𝑥))))
 
Theorembj-2inf 10733* Two formulations of the axiom of infinity (see ax-infvn 10736 and bj-omex 10737) . (Contributed by BJ, 30-Nov-2019.) (Proof modification is discouraged.)
(ω ∈ V ↔ ∃𝑥(Ind 𝑥 ∧ ∀𝑦(Ind 𝑦𝑥𝑦)))
 
6.2.5.3  The first three Peano postulates

The first three Peano postulates follow from constructive set theory (actually, from its core axioms). The proofs peano1 4335 and peano3 4337 already show this. In this section, we prove bj-peano2 10734 to complete this program. We also prove a preliminary version of the fifth Peano postulate from the core axioms.

 
Theorembj-peano2 10734 Constructive proof of peano2 4336. Temporary note: another possibility is to simply replace sucexg 4242 with bj-sucexg 10713 in the proof of peano2 4336. (Contributed by BJ, 18-Nov-2019.) (Proof modification is discouraged.)
(𝐴 ∈ ω → suc 𝐴 ∈ ω)
 
Theorempeano5set 10735* Version of peano5 4339 when ω ∩ 𝐴 is assumed to be a set, allowing a proof from the core axioms of CZF. (Contributed by BJ, 19-Nov-2019.) (Proof modification is discouraged.)
((ω ∩ 𝐴) ∈ 𝑉 → ((∅ ∈ 𝐴 ∧ ∀𝑥 ∈ ω (𝑥𝐴 → suc 𝑥𝐴)) → ω ⊆ 𝐴))
 
6.2.6  CZF: Infinity

In the absence of full separation, the axiom of infinity has to be stated more precisely, as the existence of the smallest class containing the empty set and the successor of each of its elements.

 
6.2.6.1  The set of natural numbers (finite ordinals)

In this section, we introduce the axiom of infinity in a constructive setting (ax-infvn 10736) and deduce that the class ω of finite ordinals is a set (bj-omex 10737).

 
Axiomax-infvn 10736* Axiom of infinity in a constructive setting. This asserts the existence of the special set we want (the set of natural numbers), instead of the existence of a set with some properties (ax-iinf 4329) from which one then proves, using full separation, that the wanted set exists (omex 4334). "vn" is for "von Neumann". (Contributed by BJ, 14-Nov-2019.)
𝑥(Ind 𝑥 ∧ ∀𝑦(Ind 𝑦𝑥𝑦))
 
Theorembj-omex 10737 Proof of omex 4334 from ax-infvn 10736. (Contributed by BJ, 14-Nov-2019.) (Proof modification is discouraged.)
ω ∈ V
 
6.2.6.2  Peano's fifth postulate

In this section, we give constructive proofs of two versions of Peano's fifth postulate.

 
Theorembdpeano5 10738* Bounded version of peano5 4339. (Contributed by BJ, 19-Nov-2019.) (Proof modification is discouraged.)
BOUNDED 𝐴       ((∅ ∈ 𝐴 ∧ ∀𝑥 ∈ ω (𝑥𝐴 → suc 𝑥𝐴)) → ω ⊆ 𝐴)
 
Theoremspeano5 10739* Version of peano5 4339 when 𝐴 is assumed to be a set, allowing a proof from the core axioms of CZF. (Contributed by BJ, 19-Nov-2019.) (Proof modification is discouraged.)
((𝐴𝑉 ∧ ∅ ∈ 𝐴 ∧ ∀𝑥 ∈ ω (𝑥𝐴 → suc 𝑥𝐴)) → ω ⊆ 𝐴)
 
6.2.6.3  Bounded induction and Peano's fourth postulate

In this section, we prove various versions of bounded induction from the basic axioms of CZF (in particular, without the axiom of set induction). We also prove Peano's fourth postulate. Together with the results from the previous sections, this proves from the core axioms of CZF (with infinity) that the set of finite ordinals satisfies the five Peano postulates and thus provides a model for the set of natural numbers.

 
Theoremfindset 10740* Bounded induction (principle of induction when 𝐴 is assumed to be a set) allowing a proof from basic constructive axioms. See find 4340 for a nonconstructive proof of the general case. See bdfind 10741 for a proof when 𝐴 is assumed to be bounded. (Contributed by BJ, 22-Nov-2019.) (Proof modification is discouraged.)
(𝐴𝑉 → ((𝐴 ⊆ ω ∧ ∅ ∈ 𝐴 ∧ ∀𝑥𝐴 suc 𝑥𝐴) → 𝐴 = ω))
 
Theorembdfind 10741* Bounded induction (principle of induction when 𝐴 is assumed to be bounded), proved from basic constructive axioms. See find 4340 for a nonconstructive proof of the general case. See findset 10740 for a proof when 𝐴 is assumed to be a set. (Contributed by BJ, 22-Nov-2019.) (Proof modification is discouraged.)
BOUNDED 𝐴       ((𝐴 ⊆ ω ∧ ∅ ∈ 𝐴 ∧ ∀𝑥𝐴 suc 𝑥𝐴) → 𝐴 = ω)
 
Theorembj-bdfindis 10742* Bounded induction (principle of induction for bounded formulas), using implicit substitutions (the biconditional versions of the hypotheses are implicit substitutions, and we have weakened them to implications). Constructive proof (from CZF). See finds 4341 for a proof of full induction in IZF. From this version, it is easy to prove bounded versions of finds 4341, finds2 4342, finds1 4343. (Contributed by BJ, 21-Nov-2019.) (Proof modification is discouraged.)
BOUNDED 𝜑    &   𝑥𝜓    &   𝑥𝜒    &   𝑥𝜃    &   (𝑥 = ∅ → (𝜓𝜑))    &   (𝑥 = 𝑦 → (𝜑𝜒))    &   (𝑥 = suc 𝑦 → (𝜃𝜑))       ((𝜓 ∧ ∀𝑦 ∈ ω (𝜒𝜃)) → ∀𝑥 ∈ ω 𝜑)
 
Theorembj-bdfindisg 10743* Version of bj-bdfindis 10742 using a class term in the consequent. Constructive proof (from CZF). See the comment of bj-bdfindis 10742 for explanations. (Contributed by BJ, 21-Nov-2019.) (Proof modification is discouraged.)
BOUNDED 𝜑    &   𝑥𝜓    &   𝑥𝜒    &   𝑥𝜃    &   (𝑥 = ∅ → (𝜓𝜑))    &   (𝑥 = 𝑦 → (𝜑𝜒))    &   (𝑥 = suc 𝑦 → (𝜃𝜑))    &   𝑥𝐴    &   𝑥𝜏    &   (𝑥 = 𝐴 → (𝜑𝜏))       ((𝜓 ∧ ∀𝑦 ∈ ω (𝜒𝜃)) → (𝐴 ∈ ω → 𝜏))
 
Theorembj-bdfindes 10744 Bounded induction (principle of induction for bounded formulas), using explicit substitutions. Constructive proof (from CZF). See the comment of bj-bdfindis 10742 for explanations. From this version, it is easy to prove the bounded version of findes 4344. (Contributed by BJ, 21-Nov-2019.) (Proof modification is discouraged.)
BOUNDED 𝜑       (([∅ / 𝑥]𝜑 ∧ ∀𝑥 ∈ ω (𝜑[suc 𝑥 / 𝑥]𝜑)) → ∀𝑥 ∈ ω 𝜑)
 
Theorembj-nn0suc0 10745* Constructive proof of a variant of nn0suc 4345. For a constructive proof of nn0suc 4345, see bj-nn0suc 10759. (Contributed by BJ, 19-Nov-2019.) (Proof modification is discouraged.)
(𝐴 ∈ ω → (𝐴 = ∅ ∨ ∃𝑥𝐴 𝐴 = suc 𝑥))
 
Theorembj-nntrans 10746 A natural number is a transitive set. (Contributed by BJ, 22-Nov-2019.) (Proof modification is discouraged.)
(𝐴 ∈ ω → (𝐵𝐴𝐵𝐴))
 
Theorembj-nntrans2 10747 A natural number is a transitive set. (Contributed by BJ, 22-Nov-2019.) (Proof modification is discouraged.)
(𝐴 ∈ ω → Tr 𝐴)
 
Theorembj-nnelirr 10748 A natural number does not belong to itself. Version of elirr 4284 for natural numbers, which does not require ax-setind 4280. (Contributed by BJ, 24-Nov-2019.) (Proof modification is discouraged.)
(𝐴 ∈ ω → ¬ 𝐴𝐴)
 
Theorembj-nnen2lp 10749 A version of en2lp 4297 for natural numbers, which does not require ax-setind 4280.

Note: using this theorem and bj-nnelirr 10748, one can remove dependency on ax-setind 4280 from nntri2 6096 and nndcel 6101; one can actually remove more dependencies from these. (Contributed by BJ, 28-Nov-2019.) (Proof modification is discouraged.)

((𝐴 ∈ ω ∧ 𝐵 ∈ ω) → ¬ (𝐴𝐵𝐵𝐴))
 
Theorembj-peano4 10750 Remove from peano4 4338 dependency on ax-setind 4280. Therefore, it only requires core constructive axioms (albeit more of them). (Contributed by BJ, 28-Nov-2019.) (Proof modification is discouraged.)
((𝐴 ∈ ω ∧ 𝐵 ∈ ω) → (suc 𝐴 = suc 𝐵𝐴 = 𝐵))
 
Theorembj-omtrans 10751 The set ω is transitive. A natural number is included in ω. Constructive proof of elnn 4346.

The idea is to use bounded induction with the formula 𝑥 ⊆ ω. This formula, in a logic with terms, is bounded. So in our logic without terms, we need to temporarily replace it with 𝑥𝑎 and then deduce the original claim. (Contributed by BJ, 29-Dec-2019.) (Proof modification is discouraged.)

(𝐴 ∈ ω → 𝐴 ⊆ ω)
 
Theorembj-omtrans2 10752 The set ω is transitive. (Contributed by BJ, 29-Dec-2019.) (Proof modification is discouraged.)
Tr ω
 
Theorembj-nnord 10753 A natural number is an ordinal. Constructive proof of nnord 4352. Can also be proved from bj-nnelon 10754 if the latter is proved from bj-omssonALT 10758. (Contributed by BJ, 27-Oct-2020.) (Proof modification is discouraged.)
(𝐴 ∈ ω → Ord 𝐴)
 
Theorembj-nnelon 10754 A natural number is an ordinal. Constructive proof of nnon 4350. Can also be proved from bj-omssonALT 10758. (Contributed by BJ, 27-Oct-2020.) (Proof modification is discouraged.)
(𝐴 ∈ ω → 𝐴 ∈ On)
 
Theorembj-omord 10755 The set ω is an ordinal. Constructive proof of ordom 4347. (Contributed by BJ, 29-Dec-2019.) (Proof modification is discouraged.)
Ord ω
 
Theorembj-omelon 10756 The set ω is an ordinal. Constructive proof of omelon 4349. (Contributed by BJ, 29-Dec-2019.) (Proof modification is discouraged.)
ω ∈ On
 
Theorembj-omsson 10757 Constructive proof of omsson 4353. See also bj-omssonALT 10758. (Contributed by BJ, 27-Oct-2020.) (Proof modification is discouraged.) (New usage is discouraged.
ω ⊆ On
 
Theorembj-omssonALT 10758 Alternate proof of bj-omsson 10757. (Contributed by BJ, 27-Oct-2020.) (Proof modification is discouraged.) (New usage is discouraged.)
ω ⊆ On
 
Theorembj-nn0suc 10759* Proof of (biconditional form of) nn0suc 4345 from the core axioms of CZF. See also bj-nn0sucALT 10773. As a characterization of the elements of ω, this could be labeled "elom". (Contributed by BJ, 19-Nov-2019.) (Proof modification is discouraged.)
(𝐴 ∈ ω ↔ (𝐴 = ∅ ∨ ∃𝑥 ∈ ω 𝐴 = suc 𝑥))
 
6.2.7  CZF: Set induction

In this section, we add the axiom of set induction to the core axioms of CZF.

 
6.2.7.1  Set induction

In this section, we prove some variants of the axiom of set induction.

 
Theoremsetindft 10760* Axiom of set-induction with a DV condition replaced with a non-freeness hypothesis (Contributed by BJ, 22-Nov-2019.)
(∀𝑥𝑦𝜑 → (∀𝑥(∀𝑦𝑥 [𝑦 / 𝑥]𝜑𝜑) → ∀𝑥𝜑))
 
Theoremsetindf 10761* Axiom of set-induction with a DV condition replaced with a non-freeness hypothesis (Contributed by BJ, 22-Nov-2019.)
𝑦𝜑       (∀𝑥(∀𝑦𝑥 [𝑦 / 𝑥]𝜑𝜑) → ∀𝑥𝜑)
 
Theoremsetindis 10762* Axiom of set induction using implicit substitutions. (Contributed by BJ, 22-Nov-2019.)
𝑥𝜓    &   𝑥𝜒    &   𝑦𝜑    &   𝑦𝜓    &   (𝑥 = 𝑧 → (𝜑𝜓))    &   (𝑥 = 𝑦 → (𝜒𝜑))       (∀𝑦(∀𝑧𝑦 𝜓𝜒) → ∀𝑥𝜑)
 
Axiomax-bdsetind 10763* Axiom of bounded set induction. (Contributed by BJ, 28-Nov-2019.)
BOUNDED 𝜑       (∀𝑎(∀𝑦𝑎 [𝑦 / 𝑎]𝜑𝜑) → ∀𝑎𝜑)
 
Theorembdsetindis 10764* Axiom of bounded set induction using implicit substitutions. (Contributed by BJ, 22-Nov-2019.) (Proof modification is discouraged.)
BOUNDED 𝜑    &   𝑥𝜓    &   𝑥𝜒    &   𝑦𝜑    &   𝑦𝜓    &   (𝑥 = 𝑧 → (𝜑𝜓))    &   (𝑥 = 𝑦 → (𝜒𝜑))       (∀𝑦(∀𝑧𝑦 𝜓𝜒) → ∀𝑥𝜑)
 
Theorembj-inf2vnlem1 10765* Lemma for bj-inf2vn 10769. Remark: unoptimized proof (have to use more deduction style). (Contributed by BJ, 8-Dec-2019.) (Proof modification is discouraged.)
(∀𝑥(𝑥𝐴 ↔ (𝑥 = ∅ ∨ ∃𝑦𝐴 𝑥 = suc 𝑦)) → Ind 𝐴)
 
Theorembj-inf2vnlem2 10766* Lemma for bj-inf2vnlem3 10767 and bj-inf2vnlem4 10768. Remark: unoptimized proof (have to use more deduction style). (Contributed by BJ, 8-Dec-2019.) (Proof modification is discouraged.)
(∀𝑥𝐴 (𝑥 = ∅ ∨ ∃𝑦𝐴 𝑥 = suc 𝑦) → (Ind 𝑍 → ∀𝑢(∀𝑡𝑢 (𝑡𝐴𝑡𝑍) → (𝑢𝐴𝑢𝑍))))
 
Theorembj-inf2vnlem3 10767* Lemma for bj-inf2vn 10769. (Contributed by BJ, 8-Dec-2019.) (Proof modification is discouraged.)
BOUNDED 𝐴    &   BOUNDED 𝑍       (∀𝑥𝐴 (𝑥 = ∅ ∨ ∃𝑦𝐴 𝑥 = suc 𝑦) → (Ind 𝑍𝐴𝑍))
 
Theorembj-inf2vnlem4 10768* Lemma for bj-inf2vn2 10770. (Contributed by BJ, 8-Dec-2019.) (Proof modification is discouraged.)
(∀𝑥𝐴 (𝑥 = ∅ ∨ ∃𝑦𝐴 𝑥 = suc 𝑦) → (Ind 𝑍𝐴𝑍))
 
Theorembj-inf2vn 10769* A sufficient condition for ω to be a set. See bj-inf2vn2 10770 for the unbounded version from full set induction. (Contributed by BJ, 8-Dec-2019.) (Proof modification is discouraged.)
BOUNDED 𝐴       (𝐴𝑉 → (∀𝑥(𝑥𝐴 ↔ (𝑥 = ∅ ∨ ∃𝑦𝐴 𝑥 = suc 𝑦)) → 𝐴 = ω))
 
Theorembj-inf2vn2 10770* A sufficient condition for ω to be a set; unbounded version of bj-inf2vn 10769. (Contributed by BJ, 8-Dec-2019.) (Proof modification is discouraged.)
(𝐴𝑉 → (∀𝑥(𝑥𝐴 ↔ (𝑥 = ∅ ∨ ∃𝑦𝐴 𝑥 = suc 𝑦)) → 𝐴 = ω))
 
Axiomax-inf2 10771* Another axiom of infinity in a constructive setting (see ax-infvn 10736). (Contributed by BJ, 14-Nov-2019.) (New usage is discouraged.)
𝑎𝑥(𝑥𝑎 ↔ (𝑥 = ∅ ∨ ∃𝑦𝑎 𝑥 = suc 𝑦))
 
Theorembj-omex2 10772 Using bounded set induction and the strong axiom of infinity, ω is a set, that is, we recover ax-infvn 10736 (see bj-2inf 10733 for the equivalence of the latter with bj-omex 10737). (Contributed by BJ, 8-Dec-2019.) (Proof modification is discouraged.) (New usage is discouraged.)
ω ∈ V
 
Theorembj-nn0sucALT 10773* Alternate proof of bj-nn0suc 10759, also constructive but from ax-inf2 10771, hence requiring ax-bdsetind 10763. (Contributed by BJ, 8-Dec-2019.) (Proof modification is discouraged.) (New usage is discouraged.)
(𝐴 ∈ ω ↔ (𝐴 = ∅ ∨ ∃𝑥 ∈ ω 𝐴 = suc 𝑥))
 
6.2.7.2  Full induction

In this section, using the axiom of set induction, we prove full induction on the set of natural numbers.

 
Theorembj-findis 10774* Principle of induction, using implicit substitutions (the biconditional versions of the hypotheses are implicit substitutions, and we have weakened them to implications). Constructive proof (from CZF). See bj-bdfindis 10742 for a bounded version not requiring ax-setind 4280. See finds 4341 for a proof in IZF. From this version, it is easy to prove of finds 4341, finds2 4342, finds1 4343. (Contributed by BJ, 22-Dec-2019.) (Proof modification is discouraged.)
𝑥𝜓    &   𝑥𝜒    &   𝑥𝜃    &   (𝑥 = ∅ → (𝜓𝜑))    &   (𝑥 = 𝑦 → (𝜑𝜒))    &   (𝑥 = suc 𝑦 → (𝜃𝜑))       ((𝜓 ∧ ∀𝑦 ∈ ω (𝜒𝜃)) → ∀𝑥 ∈ ω 𝜑)
 
Theorembj-findisg 10775* Version of bj-findis 10774 using a class term in the consequent. Constructive proof (from CZF). See the comment of bj-findis 10774 for explanations. (Contributed by BJ, 21-Nov-2019.) (Proof modification is discouraged.)
𝑥𝜓    &   𝑥𝜒    &   𝑥𝜃    &   (𝑥 = ∅ → (𝜓𝜑))    &   (𝑥 = 𝑦 → (𝜑𝜒))    &   (𝑥 = suc 𝑦 → (𝜃𝜑))    &   𝑥𝐴    &   𝑥𝜏    &   (𝑥 = 𝐴 → (𝜑𝜏))       ((𝜓 ∧ ∀𝑦 ∈ ω (𝜒𝜃)) → (𝐴 ∈ ω → 𝜏))
 
Theorembj-findes 10776 Principle of induction, using explicit substitutions. Constructive proof (from CZF). See the comment of bj-findis 10774 for explanations. From this version, it is easy to prove findes 4344. (Contributed by BJ, 21-Nov-2019.) (Proof modification is discouraged.)
(([∅ / 𝑥]𝜑 ∧ ∀𝑥 ∈ ω (𝜑[suc 𝑥 / 𝑥]𝜑)) → ∀𝑥 ∈ ω 𝜑)
 
6.2.8  CZF: Strong collection

In this section, we state the axiom scheme of strong collection, which is part of CZF set theory.

 
Axiomax-strcoll 10777* Axiom scheme of strong collection. It is stated with all possible disjoint variable conditions, to show that this weak form is sufficient. (Contributed by BJ, 5-Oct-2019.)
𝑎(∀𝑥𝑎𝑦𝜑 → ∃𝑏𝑦(𝑦𝑏 ↔ ∃𝑥𝑎 𝜑))
 
Theoremstrcoll2 10778* Version of ax-strcoll 10777 with one DV condition removed and without initial universal quantifier. (Contributed by BJ, 5-Oct-2019.)
(∀𝑥𝑎𝑦𝜑 → ∃𝑏𝑦(𝑦𝑏 ↔ ∃𝑥𝑎 𝜑))
 
Theoremstrcollnft 10779* Closed form of strcollnf 10780. Version of ax-strcoll 10777 with one DV condition removed, the other DV condition replaced by a non-freeness antecedent, and without initial universal quantifier. (Contributed by BJ, 21-Oct-2019.)
(∀𝑥𝑦𝑏𝜑 → (∀𝑥𝑎𝑦𝜑 → ∃𝑏𝑦(𝑦𝑏 ↔ ∃𝑥𝑎 𝜑)))
 
Theoremstrcollnf 10780* Version of ax-strcoll 10777 with one DV condition removed, the other DV condition replaced by a non-freeness hypothesis, and without initial universal quantifier. (Contributed by BJ, 21-Oct-2019.)
𝑏𝜑       (∀𝑥𝑎𝑦𝜑 → ∃𝑏𝑦(𝑦𝑏 ↔ ∃𝑥𝑎 𝜑))
 
TheoremstrcollnfALT 10781* Alternate proof of strcollnf 10780, not using strcollnft 10779. (Contributed by BJ, 5-Oct-2019.) (Proof modification is discouraged.) (New usage is discouraged.)
𝑏𝜑       (∀𝑥𝑎𝑦𝜑 → ∃𝑏𝑦(𝑦𝑏 ↔ ∃𝑥𝑎 𝜑))
 
6.2.9  CZF: Subset collection

In this section, we state the axiom scheme of subset collection, which is part of CZF set theory.

 
Axiomax-sscoll 10782* Axiom scheme of subset collection. It is stated with all possible disjoint variable conditions, to show that this weak form is sufficient. (Contributed by BJ, 5-Oct-2019.)
𝑎𝑏𝑐𝑧(∀𝑥𝑎𝑦𝑏 𝜑 → ∃𝑑𝑐𝑦(𝑦𝑑 ↔ ∃𝑥𝑎 𝜑))
 
Theoremsscoll2 10783* Version of ax-sscoll 10782 with two DV conditions removed and without initial universal quantifiers. (Contributed by BJ, 5-Oct-2019.)
𝑐𝑧(∀𝑥𝑎𝑦𝑏 𝜑 → ∃𝑑𝑐𝑦(𝑦𝑑 ↔ ∃𝑥𝑎 𝜑))
 
6.2.10  Real numbers
 
Axiomax-ddkcomp 10784 Axiom of Dedekind completeness for Dedekind real numbers: every inhabited upper-bounded located set of reals has a real upper bound. Ideally, this axiom should be "proved" as "axddkcomp" for the real numbers constructed from IZF, and then the axiom ax-ddkcomp 10784 should be used in place of construction specific results. In particular, axcaucvg 7066 should be proved from it. (Contributed by BJ, 24-Oct-2021.)
(((𝐴 ⊆ ℝ ∧ ∃𝑥 𝑥𝐴) ∧ ∃𝑥 ∈ ℝ ∀𝑦𝐴 𝑦 < 𝑥 ∧ ∀𝑥 ∈ ℝ ∀𝑦 ∈ ℝ (𝑥 < 𝑦 → (∃𝑧𝐴 𝑥 < 𝑧 ∨ ∀𝑧𝐴 𝑧 < 𝑦))) → ∃𝑥 ∈ ℝ (∀𝑦𝐴 𝑦𝑥 ∧ ((𝐵𝑅 ∧ ∀𝑦𝐴 𝑦𝐵) → 𝑥𝐵)))
 
6.3  Mathbox for Jim Kingdon
 
Theoremqdencn 10785* The set of complex numbers whose real and imaginary parts are rational is dense in the complex plane. This is a two dimensional analogue to qdenre 10088 (and also would hold for ℝ × ℝ with the usual metric; this is not about complex numbers in particular). (Contributed by Jim Kingdon, 18-Oct-2021.)
𝑄 = {𝑧 ∈ ℂ ∣ ((ℜ‘𝑧) ∈ ℚ ∧ (ℑ‘𝑧) ∈ ℚ)}       ((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℝ+) → ∃𝑥𝑄 (abs‘(𝑥𝐴)) < 𝐵)
 
6.4  Mathbox for Mykola Mostovenko
 
Theoremax1hfs 10786 Heyting's formal system Axiom #1 from [Heyting] p. 127. (Contributed by MM, 11-Aug-2018.)
(𝜑 → (𝜑𝜑))
 
6.5  Mathbox for David A. Wheeler
 
6.5.1  Allsome quantifier

These are definitions and proofs involving an experimental "allsome" quantifier (aka "all some").

In informal language, statements like "All Martians are green" imply that there is at least one Martian. But it's easy to mistranslate informal language into formal notations because similar statements like 𝑥𝜑𝜓 do not imply that 𝜑 is ever true, leading to vacuous truths. Some systems include a mechanism to counter this, e.g., PVS allows types to be appended with "+" to declare that they are nonempty. This section presents a different solution to the same problem.

The "allsome" quantifier expressly includes the notion of both "all" and "there exists at least one" (aka some), and is defined to make it easier to more directly express both notions. The hope is that if a quantifier more directly expresses this concept, it will be used instead and reduce the risk of creating formal expressions that look okay but in fact are mistranslations. The term "allsome" was chosen because it's short, easy to say, and clearly hints at the two concepts it combines.

I do not expect this to be used much in metamath, because in metamath there's a general policy of avoiding the use of new definitions unless there are very strong reasons to do so. Instead, my goal is to rigorously define this quantifier and demonstrate a few basic properties of it.

The syntax allows two forms that look like they would be problematic, but they are fine. When applied to a top-level implication we allow ∀!𝑥(𝜑𝜓), and when restricted (applied to a class) we allow ∀!𝑥𝐴𝜑. The first symbol after the setvar variable must always be if it is the form applied to a class, and since cannot begin a wff, it is unambiguous. The looks like it would be a problem because 𝜑 or 𝜓 might include implications, but any implication arrow within any wff must be surrounded by parentheses, so only the implication arrow of ∀! can follow the wff. The implication syntax would work fine without the parentheses, but I added the parentheses because it makes things clearer inside larger complex expressions, and it's also more consistent with the rest of the syntax.

For more, see "The Allsome Quantifier" by David A. Wheeler at https://dwheeler.com/essays/allsome.html I hope that others will eventually agree that allsome is awesome.

 
Syntaxwalsi 10787 Extend wff definition to include "all some" applied to a top-level implication, which means 𝜓 is true whenever 𝜑 is true, and there is at least least one 𝑥 where 𝜑 is true. (Contributed by David A. Wheeler, 20-Oct-2018.)
wff ∀!𝑥(𝜑𝜓)
 
Syntaxwalsc 10788 Extend wff definition to include "all some" applied to a class, which means 𝜑 is true for all 𝑥 in 𝐴, and there is at least one 𝑥 in 𝐴. (Contributed by David A. Wheeler, 20-Oct-2018.)
wff ∀!𝑥𝐴𝜑
 
Definitiondf-alsi 10789 Define "all some" applied to a top-level implication, which means 𝜓 is true whenever 𝜑 is true and there is at least one 𝑥 where 𝜑 is true. (Contributed by David A. Wheeler, 20-Oct-2018.)
(∀!𝑥(𝜑𝜓) ↔ (∀𝑥(𝜑𝜓) ∧ ∃𝑥𝜑))
 
Definitiondf-alsc 10790 Define "all some" applied to a class, which means 𝜑 is true for all 𝑥 in 𝐴 and there is at least one 𝑥 in 𝐴. (Contributed by David A. Wheeler, 20-Oct-2018.)
(∀!𝑥𝐴𝜑 ↔ (∀𝑥𝐴 𝜑 ∧ ∃𝑥 𝑥𝐴))
 
Theoremalsconv 10791 There is an equivalence between the two "all some" forms. (Contributed by David A. Wheeler, 22-Oct-2018.)
(∀!𝑥(𝑥𝐴𝜑) ↔ ∀!𝑥𝐴𝜑)
 
Theoremalsi1d 10792 Deduction rule: Given "all some" applied to a top-level inference, you can extract the "for all" part. (Contributed by David A. Wheeler, 20-Oct-2018.)
(𝜑 → ∀!𝑥(𝜓𝜒))       (𝜑 → ∀𝑥(𝜓𝜒))
 
Theoremalsi2d 10793 Deduction rule: Given "all some" applied to a top-level inference, you can extract the "exists" part. (Contributed by David A. Wheeler, 20-Oct-2018.)
(𝜑 → ∀!𝑥(𝜓𝜒))       (𝜑 → ∃𝑥𝜓)
 
Theoremalsc1d 10794 Deduction rule: Given "all some" applied to a class, you can extract the "for all" part. (Contributed by David A. Wheeler, 20-Oct-2018.)
(𝜑 → ∀!𝑥𝐴𝜓)       (𝜑 → ∀𝑥𝐴 𝜓)
 
Theoremalsc2d 10795 Deduction rule: Given "all some" applied to a class, you can extract the "there exists" part. (Contributed by David A. Wheeler, 20-Oct-2018.)
(𝜑 → ∀!𝑥𝐴𝜓)       (𝜑 → ∃𝑥 𝑥𝐴)
    < Previous  Wrap >

Page List
Jump to page: Contents  1 1-100 2 101-200 3 201-300 4 301-400 5 401-500 6 501-600 7 601-700 8 701-800 9 801-900 10 901-1000 11 1001-1100 12 1101-1200 13 1201-1300 14 1301-1400 15 1401-1500 16 1501-1600 17 1601-1700 18 1701-1800 19 1801-1900 20 1901-2000 21 2001-2100 22 2101-2200 23 2201-2300 24 2301-2400 25 2401-2500 26 2501-2600 27 2601-2700 28 2701-2800 29 2801-2900 30 2901-3000 31 3001-3100 32 3101-3200 33 3201-3300 34 3301-3400 35 3401-3500 36 3501-3600 37 3601-3700 38 3701-3800 39 3801-3900 40 3901-4000 41 4001-4100 42 4101-4200 43 4201-4300 44 4301-4400 45 4401-4500 46 4501-4600 47 4601-4700 48 4701-4800 49 4801-4900 50 4901-5000 51 5001-5100 52 5101-5200 53 5201-5300 54 5301-5400 55 5401-5500 56 5501-5600 57 5601-5700 58 5701-5800 59 5801-5900 60 5901-6000 61 6001-6100 62 6101-6200 63 6201-6300 64 6301-6400 65 6401-6500 66 6501-6600 67 6601-6700 68 6701-6800 69 6801-6900 70 6901-7000 71 7001-7100 72 7101-7200 73 7201-7300 74 7301-7400 75 7401-7500 76 7501-7600 77 7601-7700 78 7701-7800 79 7801-7900 80 7901-8000 81 8001-8100 82 8101-8200 83 8201-8300 84 8301-8400 85 8401-8500 86 8501-8600 87 8601-8700 88 8701-8800 89 8801-8900 90 8901-9000 91 9001-9100 92 9101-9200 93 9201-9300 94 9301-9400 95 9401-9500 96 9501-9600 97 9601-9700 98 9701-9800 99 9801-9900 100 9901-10000 101 10001-10100 102 10101-10200 103 10201-10300 104 10301-10400 105 10401-10500 106 10501-10600 107 10601-10700 108 10701-10795
  Copyright terms: Public domain < Previous  Wrap >