HomeHome Intuitionistic Logic Explorer
Theorem List (p. 106 of 108)
< Previous  Next >
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 - 10501-10600   *Has distinct variable group(s)
TypeLabelDescription
Statement
 
Theoremisprm4 10501* The predicate "is a prime number". A prime number is an integer greater than or equal to 2 whose only divisor greater than or equal to 2 is itself. (Contributed by Paul Chapman, 26-Oct-2012.)
(𝑃 ∈ ℙ ↔ (𝑃 ∈ (ℤ‘2) ∧ ∀𝑧 ∈ (ℤ‘2)(𝑧𝑃𝑧 = 𝑃)))
 
Theoremprmind2 10502* A variation on prmind 10503 assuming complete induction for primes. (Contributed by Mario Carneiro, 20-Jun-2015.)
(𝑥 = 1 → (𝜑𝜓))    &   (𝑥 = 𝑦 → (𝜑𝜒))    &   (𝑥 = 𝑧 → (𝜑𝜃))    &   (𝑥 = (𝑦 · 𝑧) → (𝜑𝜏))    &   (𝑥 = 𝐴 → (𝜑𝜂))    &   𝜓    &   ((𝑥 ∈ ℙ ∧ ∀𝑦 ∈ (1...(𝑥 − 1))𝜒) → 𝜑)    &   ((𝑦 ∈ (ℤ‘2) ∧ 𝑧 ∈ (ℤ‘2)) → ((𝜒𝜃) → 𝜏))       (𝐴 ∈ ℕ → 𝜂)
 
Theoremprmind 10503* Perform induction over the multiplicative structure of . If a property 𝜑(𝑥) holds for the primes and 1 and is preserved under multiplication, then it holds for every positive integer. (Contributed by Mario Carneiro, 20-Jun-2015.)
(𝑥 = 1 → (𝜑𝜓))    &   (𝑥 = 𝑦 → (𝜑𝜒))    &   (𝑥 = 𝑧 → (𝜑𝜃))    &   (𝑥 = (𝑦 · 𝑧) → (𝜑𝜏))    &   (𝑥 = 𝐴 → (𝜑𝜂))    &   𝜓    &   (𝑥 ∈ ℙ → 𝜑)    &   ((𝑦 ∈ (ℤ‘2) ∧ 𝑧 ∈ (ℤ‘2)) → ((𝜒𝜃) → 𝜏))       (𝐴 ∈ ℕ → 𝜂)
 
Theoremdvdsprime 10504 If 𝑀 divides a prime, then 𝑀 is either the prime or one. (Contributed by Scott Fenton, 8-Apr-2014.)
((𝑃 ∈ ℙ ∧ 𝑀 ∈ ℕ) → (𝑀𝑃 ↔ (𝑀 = 𝑃𝑀 = 1)))
 
Theoremnprm 10505 A product of two integers greater than one is composite. (Contributed by Mario Carneiro, 20-Jun-2015.)
((𝐴 ∈ (ℤ‘2) ∧ 𝐵 ∈ (ℤ‘2)) → ¬ (𝐴 · 𝐵) ∈ ℙ)
 
Theoremnprmi 10506 An inference for compositeness. (Contributed by Mario Carneiro, 18-Feb-2014.) (Revised by Mario Carneiro, 20-Jun-2015.)
𝐴 ∈ ℕ    &   𝐵 ∈ ℕ    &   1 < 𝐴    &   1 < 𝐵    &   (𝐴 · 𝐵) = 𝑁        ¬ 𝑁 ∈ ℙ
 
Theoremdvdsnprmd 10507 If a number is divisible by an integer greater than 1 and less then the number, the number is not prime. (Contributed by AV, 24-Jul-2021.)
(𝜑 → 1 < 𝐴)    &   (𝜑𝐴 < 𝑁)    &   (𝜑𝐴𝑁)       (𝜑 → ¬ 𝑁 ∈ ℙ)
 
Theoremprm2orodd 10508 A prime number is either 2 or odd. (Contributed by AV, 19-Jun-2021.)
(𝑃 ∈ ℙ → (𝑃 = 2 ∨ ¬ 2 ∥ 𝑃))
 
Theorem2prm 10509 2 is a prime number. (Contributed by Paul Chapman, 22-Jun-2011.) (Proof shortened by Fan Zheng, 16-Jun-2016.)
2 ∈ ℙ
 
Theorem3prm 10510 3 is a prime number. (Contributed by Paul Chapman, 22-Jun-2011.)
3 ∈ ℙ
 
Theorem4nprm 10511 4 is not a prime number. (Contributed by Paul Chapman, 22-Jun-2011.) (Proof shortened by Mario Carneiro, 18-Feb-2014.)
¬ 4 ∈ ℙ
 
Theoremprmuz2 10512 A prime number is an integer greater than or equal to 2. (Contributed by Paul Chapman, 17-Nov-2012.)
(𝑃 ∈ ℙ → 𝑃 ∈ (ℤ‘2))
 
Theoremprmgt1 10513 A prime number is an integer greater than 1. (Contributed by Alexander van der Vekens, 17-May-2018.)
(𝑃 ∈ ℙ → 1 < 𝑃)
 
Theoremprmm2nn0 10514 Subtracting 2 from a prime number results in a nonnegative integer. (Contributed by Alexander van der Vekens, 30-Aug-2018.)
(𝑃 ∈ ℙ → (𝑃 − 2) ∈ ℕ0)
 
Theoremoddprmgt2 10515 An odd prime is greater than 2. (Contributed by AV, 20-Aug-2021.)
(𝑃 ∈ (ℙ ∖ {2}) → 2 < 𝑃)
 
Theoremoddprmge3 10516 An odd prime is greater than or equal to 3. (Contributed by Alexander van der Vekens, 7-Oct-2018.) (Revised by AV, 20-Aug-2021.)
(𝑃 ∈ (ℙ ∖ {2}) → 𝑃 ∈ (ℤ‘3))
 
Theoremsqnprm 10517 A square is never prime. (Contributed by Mario Carneiro, 20-Jun-2015.)
(𝐴 ∈ ℤ → ¬ (𝐴↑2) ∈ ℙ)
 
Theoremdvdsprm 10518 An integer greater than or equal to 2 divides a prime number iff it is equal to it. (Contributed by Paul Chapman, 26-Oct-2012.)
((𝑁 ∈ (ℤ‘2) ∧ 𝑃 ∈ ℙ) → (𝑁𝑃𝑁 = 𝑃))
 
Theoremexprmfct 10519* Every integer greater than or equal to 2 has a prime factor. (Contributed by Paul Chapman, 26-Oct-2012.) (Proof shortened by Mario Carneiro, 20-Jun-2015.)
(𝑁 ∈ (ℤ‘2) → ∃𝑝 ∈ ℙ 𝑝𝑁)
 
Theoremprmdvdsfz 10520* Each integer greater than 1 and less then or equal to a fixed number is divisible by a prime less then or equal to this fixed number. (Contributed by AV, 15-Aug-2020.)
((𝑁 ∈ ℕ ∧ 𝐼 ∈ (2...𝑁)) → ∃𝑝 ∈ ℙ (𝑝𝑁𝑝𝐼))
 
Theoremnprmdvds1 10521 No prime number divides 1. (Contributed by Paul Chapman, 17-Nov-2012.) (Proof shortened by Mario Carneiro, 2-Jul-2015.)
(𝑃 ∈ ℙ → ¬ 𝑃 ∥ 1)
 
Theoremdivgcdodd 10522 Either 𝐴 / (𝐴 gcd 𝐵) is odd or 𝐵 / (𝐴 gcd 𝐵) is odd. (Contributed by Scott Fenton, 19-Apr-2014.)
((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → (¬ 2 ∥ (𝐴 / (𝐴 gcd 𝐵)) ∨ ¬ 2 ∥ (𝐵 / (𝐴 gcd 𝐵))))
 
4.2.2  Coprimality and Euclid's lemma (cont.)

This section is about coprimality with respect to primes, and a special version of Euclid's lemma for primes is provided, see euclemma 10525.

 
Theoremcoprm 10523 A prime number either divides an integer or is coprime to it, but not both. Theorem 1.8 in [ApostolNT] p. 17. (Contributed by Paul Chapman, 22-Jun-2011.)
((𝑃 ∈ ℙ ∧ 𝑁 ∈ ℤ) → (¬ 𝑃𝑁 ↔ (𝑃 gcd 𝑁) = 1))
 
Theoremprmrp 10524 Unequal prime numbers are relatively prime. (Contributed by Mario Carneiro, 23-Feb-2014.)
((𝑃 ∈ ℙ ∧ 𝑄 ∈ ℙ) → ((𝑃 gcd 𝑄) = 1 ↔ 𝑃𝑄))
 
Theoremeuclemma 10525 Euclid's lemma. A prime number divides the product of two integers iff it divides at least one of them. Theorem 1.9 in [ApostolNT] p. 17. (Contributed by Paul Chapman, 17-Nov-2012.)
((𝑃 ∈ ℙ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝑃 ∥ (𝑀 · 𝑁) ↔ (𝑃𝑀𝑃𝑁)))
 
Theoremisprm6 10526* A number is prime iff it satisfies Euclid's lemma euclemma 10525. (Contributed by Mario Carneiro, 6-Sep-2015.)
(𝑃 ∈ ℙ ↔ (𝑃 ∈ (ℤ‘2) ∧ ∀𝑥 ∈ ℤ ∀𝑦 ∈ ℤ (𝑃 ∥ (𝑥 · 𝑦) → (𝑃𝑥𝑃𝑦))))
 
Theoremprmdvdsexp 10527 A prime divides a positive power of an integer iff it divides the integer. (Contributed by Mario Carneiro, 24-Feb-2014.) (Revised by Mario Carneiro, 17-Jul-2014.)
((𝑃 ∈ ℙ ∧ 𝐴 ∈ ℤ ∧ 𝑁 ∈ ℕ) → (𝑃 ∥ (𝐴𝑁) ↔ 𝑃𝐴))
 
Theoremprmdvdsexpb 10528 A prime divides a positive power of another iff they are equal. (Contributed by Paul Chapman, 30-Nov-2012.) (Revised by Mario Carneiro, 24-Feb-2014.)
((𝑃 ∈ ℙ ∧ 𝑄 ∈ ℙ ∧ 𝑁 ∈ ℕ) → (𝑃 ∥ (𝑄𝑁) ↔ 𝑃 = 𝑄))
 
Theoremprmdvdsexpr 10529 If a prime divides a nonnegative power of another, then they are equal. (Contributed by Mario Carneiro, 16-Jan-2015.)
((𝑃 ∈ ℙ ∧ 𝑄 ∈ ℙ ∧ 𝑁 ∈ ℕ0) → (𝑃 ∥ (𝑄𝑁) → 𝑃 = 𝑄))
 
Theoremprmexpb 10530 Two positive prime powers are equal iff the primes and the powers are equal. (Contributed by Paul Chapman, 30-Nov-2012.)
(((𝑃 ∈ ℙ ∧ 𝑄 ∈ ℙ) ∧ (𝑀 ∈ ℕ ∧ 𝑁 ∈ ℕ)) → ((𝑃𝑀) = (𝑄𝑁) ↔ (𝑃 = 𝑄𝑀 = 𝑁)))
 
Theoremprmfac1 10531 The factorial of a number only contains primes less than the base. (Contributed by Mario Carneiro, 6-Mar-2014.)
((𝑁 ∈ ℕ0𝑃 ∈ ℙ ∧ 𝑃 ∥ (!‘𝑁)) → 𝑃𝑁)
 
Theoremrpexp 10532 If two numbers 𝐴 and 𝐵 are relatively prime, then they are still relatively prime if raised to a power. (Contributed by Mario Carneiro, 24-Feb-2014.)
((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ ∧ 𝑁 ∈ ℕ) → (((𝐴𝑁) gcd 𝐵) = 1 ↔ (𝐴 gcd 𝐵) = 1))
 
Theoremrpexp1i 10533 Relative primality passes to asymmetric powers. (Contributed by Stefan O'Rear, 27-Sep-2014.)
((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ ∧ 𝑀 ∈ ℕ0) → ((𝐴 gcd 𝐵) = 1 → ((𝐴𝑀) gcd 𝐵) = 1))
 
Theoremrpexp12i 10534 Relative primality passes to symmetric powers. (Contributed by Stefan O'Rear, 27-Sep-2014.)
((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ ∧ (𝑀 ∈ ℕ0𝑁 ∈ ℕ0)) → ((𝐴 gcd 𝐵) = 1 → ((𝐴𝑀) gcd (𝐵𝑁)) = 1))
 
Theoremprmndvdsfaclt 10535 A prime number does not divide the factorial of a nonnegative integer less than the prime number. (Contributed by AV, 13-Jul-2021.)
((𝑃 ∈ ℙ ∧ 𝑁 ∈ ℕ0) → (𝑁 < 𝑃 → ¬ 𝑃 ∥ (!‘𝑁)))
 
Theoremcncongrprm 10536 Corollary 2 of Cancellability of Congruences: Two products with a common factor are congruent modulo a prime number not dividing the common factor iff the other factors are congruent modulo the prime number. (Contributed by AV, 13-Jul-2021.)
(((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ ∧ 𝐶 ∈ ℤ) ∧ (𝑃 ∈ ℙ ∧ ¬ 𝑃𝐶)) → (((𝐴 · 𝐶) mod 𝑃) = ((𝐵 · 𝐶) mod 𝑃) ↔ (𝐴 mod 𝑃) = (𝐵 mod 𝑃)))
 
Theoremisevengcd2 10537 The predicate "is an even number". An even number and 2 have 2 as greatest common divisor. (Contributed by AV, 1-Jul-2020.) (Revised by AV, 8-Aug-2021.)
(𝑍 ∈ ℤ → (2 ∥ 𝑍 ↔ (2 gcd 𝑍) = 2))
 
Theoremisoddgcd1 10538 The predicate "is an odd number". An odd number and 2 have 1 as greatest common divisor. (Contributed by AV, 1-Jul-2020.) (Revised by AV, 8-Aug-2021.)
(𝑍 ∈ ℤ → (¬ 2 ∥ 𝑍 ↔ (2 gcd 𝑍) = 1))
 
Theorem3lcm2e6 10539 The least common multiple of three and two is six. The operands are unequal primes and thus coprime, so the result is (the absolute value of) their product. (Contributed by Steve Rodriguez, 20-Jan-2020.) (Proof shortened by AV, 27-Aug-2020.)
(3 lcm 2) = 6
 
4.2.3  Non-rationality of square root of 2
 
Theoremsqrt2irrlem 10540 Lemma for sqrt2irr 10541. This is the core of the proof: - if 𝐴 / 𝐵 = √(2), then 𝐴 and 𝐵 are even, so 𝐴 / 2 and 𝐵 / 2 are smaller representatives, which is absurd by the method of infinite descent (here implemented by strong induction). (Contributed by NM, 20-Aug-2001.) (Revised by Mario Carneiro, 12-Sep-2015.)
(𝜑𝐴 ∈ ℤ)    &   (𝜑𝐵 ∈ ℕ)    &   (𝜑 → (√‘2) = (𝐴 / 𝐵))       (𝜑 → ((𝐴 / 2) ∈ ℤ ∧ (𝐵 / 2) ∈ ℕ))
 
Theoremsqrt2irr 10541 The square root of 2 is not rational. That is, for any rational number, (√‘2) does not equal it. However, if we were to say "the square root of 2 is irrational" that would mean something stronger: "for any rational number, (√‘2) is apart from it" (the two statements are equivalent given excluded middle). See sqrt2irrap 10558 for the proof that the square root of two is irrational.

The proof's core is proven in sqrt2irrlem 10540, which shows that if 𝐴 / 𝐵 = √(2), then 𝐴 and 𝐵 are even, so 𝐴 / 2 and 𝐵 / 2 are smaller representatives, which is absurd. (Contributed by NM, 8-Jan-2002.) (Proof shortened by Mario Carneiro, 12-Sep-2015.)

(√‘2) ∉ ℚ
 
Theoremsqrt2re 10542 The square root of 2 exists and is a real number. (Contributed by NM, 3-Dec-2004.)
(√‘2) ∈ ℝ
 
Theorempw2dvdslemn 10543* Lemma for pw2dvds 10544. If a natural number has some power of two which does not divide it, there is a highest power of two which does divide it. (Contributed by Jim Kingdon, 14-Nov-2021.)
((𝑁 ∈ ℕ ∧ 𝐴 ∈ ℕ ∧ ¬ (2↑𝐴) ∥ 𝑁) → ∃𝑚 ∈ ℕ0 ((2↑𝑚) ∥ 𝑁 ∧ ¬ (2↑(𝑚 + 1)) ∥ 𝑁))
 
Theorempw2dvds 10544* A natural number has a highest power of two which divides it. (Contributed by Jim Kingdon, 14-Nov-2021.)
(𝑁 ∈ ℕ → ∃𝑚 ∈ ℕ0 ((2↑𝑚) ∥ 𝑁 ∧ ¬ (2↑(𝑚 + 1)) ∥ 𝑁))
 
Theorempw2dvdseulemle 10545 Lemma for pw2dvdseu 10546. Powers of two which do and do not divide a natural number. (Contributed by Jim Kingdon, 17-Nov-2021.)
(𝜑𝑁 ∈ ℕ)    &   (𝜑𝐴 ∈ ℕ0)    &   (𝜑𝐵 ∈ ℕ0)    &   (𝜑 → (2↑𝐴) ∥ 𝑁)    &   (𝜑 → ¬ (2↑(𝐵 + 1)) ∥ 𝑁)       (𝜑𝐴𝐵)
 
Theorempw2dvdseu 10546* A natural number has a unique highest power of two which divides it. (Contributed by Jim Kingdon, 16-Nov-2021.)
(𝑁 ∈ ℕ → ∃!𝑚 ∈ ℕ0 ((2↑𝑚) ∥ 𝑁 ∧ ¬ (2↑(𝑚 + 1)) ∥ 𝑁))
 
Theoremoddpwdclemxy 10547* Lemma for oddpwdc 10552. Another way of stating that decomposing a natural number into a power of two and an odd number is unique. (Contributed by Jim Kingdon, 16-Nov-2021.)
((((𝑋 ∈ ℕ ∧ ¬ 2 ∥ 𝑋) ∧ 𝑌 ∈ ℕ0) ∧ 𝐴 = ((2↑𝑌) · 𝑋)) → (𝑋 = (𝐴 / (2↑(𝑧 ∈ ℕ0 ((2↑𝑧) ∥ 𝐴 ∧ ¬ (2↑(𝑧 + 1)) ∥ 𝐴)))) ∧ 𝑌 = (𝑧 ∈ ℕ0 ((2↑𝑧) ∥ 𝐴 ∧ ¬ (2↑(𝑧 + 1)) ∥ 𝐴))))
 
Theoremoddpwdclemdvds 10548* Lemma for oddpwdc 10552. A natural number is divisible by the highest power of two which divides it. (Contributed by Jim Kingdon, 17-Nov-2021.)
(𝐴 ∈ ℕ → (2↑(𝑧 ∈ ℕ0 ((2↑𝑧) ∥ 𝐴 ∧ ¬ (2↑(𝑧 + 1)) ∥ 𝐴))) ∥ 𝐴)
 
Theoremoddpwdclemndvds 10549* Lemma for oddpwdc 10552. A natural number is not divisible by one more than the highest power of two which divides it. (Contributed by Jim Kingdon, 17-Nov-2021.)
(𝐴 ∈ ℕ → ¬ (2↑((𝑧 ∈ ℕ0 ((2↑𝑧) ∥ 𝐴 ∧ ¬ (2↑(𝑧 + 1)) ∥ 𝐴)) + 1)) ∥ 𝐴)
 
Theoremoddpwdclemodd 10550* Lemma for oddpwdc 10552. Removing the powers of two from a natural number produces an odd number. (Contributed by Jim Kingdon, 16-Nov-2021.)
(𝐴 ∈ ℕ → ¬ 2 ∥ (𝐴 / (2↑(𝑧 ∈ ℕ0 ((2↑𝑧) ∥ 𝐴 ∧ ¬ (2↑(𝑧 + 1)) ∥ 𝐴)))))
 
Theoremoddpwdclemdc 10551* Lemma for oddpwdc 10552. Decomposing a number into odd and even parts. (Contributed by Jim Kingdon, 16-Nov-2021.)
((((𝑋 ∈ ℕ ∧ ¬ 2 ∥ 𝑋) ∧ 𝑌 ∈ ℕ0) ∧ 𝐴 = ((2↑𝑌) · 𝑋)) ↔ (𝐴 ∈ ℕ ∧ (𝑋 = (𝐴 / (2↑(𝑧 ∈ ℕ0 ((2↑𝑧) ∥ 𝐴 ∧ ¬ (2↑(𝑧 + 1)) ∥ 𝐴)))) ∧ 𝑌 = (𝑧 ∈ ℕ0 ((2↑𝑧) ∥ 𝐴 ∧ ¬ (2↑(𝑧 + 1)) ∥ 𝐴)))))
 
Theoremoddpwdc 10552* The function 𝐹 that decomposes a number into its "odd" and "even" parts, which is to say the largest power of two and largest odd divisor of a number, is a bijection from pairs of a nonnegative integer and an odd number to positive integers. (Contributed by Thierry Arnoux, 15-Aug-2017.)
𝐽 = {𝑧 ∈ ℕ ∣ ¬ 2 ∥ 𝑧}    &   𝐹 = (𝑥𝐽, 𝑦 ∈ ℕ0 ↦ ((2↑𝑦) · 𝑥))       𝐹:(𝐽 × ℕ0)–1-1-onto→ℕ
 
Theoremsqpweven 10553* The greatest power of two dividing the square of an integer is an even power of two. (Contributed by Jim Kingdon, 17-Nov-2021.)
𝐽 = {𝑧 ∈ ℕ ∣ ¬ 2 ∥ 𝑧}    &   𝐹 = (𝑥𝐽, 𝑦 ∈ ℕ0 ↦ ((2↑𝑦) · 𝑥))       (𝐴 ∈ ℕ → 2 ∥ (2nd ‘(𝐹‘(𝐴↑2))))
 
Theorem2sqpwodd 10554* The greatest power of two dividing twice the square of an integer is an odd power of two. (Contributed by Jim Kingdon, 17-Nov-2021.)
𝐽 = {𝑧 ∈ ℕ ∣ ¬ 2 ∥ 𝑧}    &   𝐹 = (𝑥𝐽, 𝑦 ∈ ℕ0 ↦ ((2↑𝑦) · 𝑥))       (𝐴 ∈ ℕ → ¬ 2 ∥ (2nd ‘(𝐹‘(2 · (𝐴↑2)))))
 
Theoremsqne2sq 10555 The square of a natural number can never be equal to two times the square of a natural number. (Contributed by Jim Kingdon, 17-Nov-2021.)
((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → (𝐴↑2) ≠ (2 · (𝐵↑2)))
 
Theoremznege1 10556 The absolute value of the difference between two unequal integers is at least one. (Contributed by Jim Kingdon, 31-Jan-2022.)
((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ ∧ 𝐴𝐵) → 1 ≤ (abs‘(𝐴𝐵)))
 
Theoremsqrt2irraplemnn 10557 Lemma for sqrt2irrap 10558. The square root of 2 is apart from a positive rational expressed as a numerator and denominator. (Contributed by Jim Kingdon, 2-Oct-2021.)
((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → (√‘2) # (𝐴 / 𝐵))
 
Theoremsqrt2irrap 10558 The square root of 2 is irrational. That is, for any rational number, (√‘2) is apart from it. In the absence of excluded middle, we can distinguish between this and "the square root of 2 is not rational" which is sqrt2irr 10541. (Contributed by Jim Kingdon, 2-Oct-2021.)
(𝑄 ∈ ℚ → (√‘2) # 𝑄)
 
PART 5  GUIDES AND MISCELLANEA
 
5.1  Guides (conventions, explanations, and examples)
 
5.1.1  Conventions

This section describes the conventions we use. However, these conventions often refer to existing mathematical practices, which are discussed in more detail in other references. The following sources lay out how mathematics is developed without the law of the excluded middle. Of course, there are a greater number of sources which assume excluded middle and most of what is in them applies here too (especially in a treatment such as ours which is built on first order logic and set theory, rather than, say, type theory). Studying how a topic is treated in the Metamath Proof Explorer and the references therein is often a good place to start (and is easy to compare with the Intuitionistic Logic Explorer). The textbooks provide a motivation for what we are doing, whereas Metamath lets you see in detail all hidden and implicit steps. Most standard theorems are accompanied by citations. Some closely followed texts include the following:

  • Axioms of propositional calculus - Stanford Encyclopedia of Philosophy or [Heyting].
  • Axioms of predicate calculus - our axioms are adapted from the ones in the Metamath Proof Explorer.
  • Theorems of propositional calculus - [Heyting].
  • Theorems of pure predicate calculus - Metamath Proof Explorer.
  • Theorems of equality and substitution - Metamath Proof Explorer.
  • Axioms of set theory - [Crosilla].
  • Development of set theory - Chapter 10 of [HoTT].
  • Construction of real and complex numbers - Chapter 11 of [HoTT]; [BauerTaylor].
  • Theorems about real numbers - [Geuvers].
 
Theoremconventions 10559 Unless there is a reason to diverge, we follow the conventions of the Metamath Proof Explorer (aka "set.mm"). This list of conventions is intended to be read in conjunction with the corresponding conventions in the Metamath Proof Explorer, and only the differences are described below.
  • Minimizing axioms and the axiom of choice. We prefer proofs that depend on fewer and/or weaker axioms, even if the proofs are longer. In particular, our choice of IZF (Intuitionistic Zermelo-Fraenkel) over CZF (Constructive Zermelo-Fraenkel, a weaker system) was just an expedient choice because IZF is easier to formalize in Metamath. You can find some development using CZF in BJ's mathbox starting at wbd 10603 (and the section header just above it). As for the axiom of choice, the full axiom of choice implies excluded middle as seen at acexmid 5531, although some authors will use countable choice or dependent choice. For example, countable choice or excluded middle is needed to show that the Cauchy reals coincide with the Dedekind reals - Corollary 11.4.3 of [HoTT], p. (varies).
  • Junk/undefined results. Much of the discussion of this topic in the Metamath Proof Explorer applies except that certain techniques are not available to us. For example, the Metamath Proof Explorer will often say "if a function is evaluated within its domain, a certain result follows; if the function is evaluated outside its domain, the same result follows. Since the function must be evaluated within its domain or outside it, the result follows unconditionally" (the use of excluded middle in this argument is perhaps obvious when stated this way). For this reason, we generally need to prove we are evaluating functions within their domains and avoid the reverse closure theorems of the Metamath Proof Explorer.
  • Bibliography references. The bibliography for the Intuitionistic Logic Explorer is separate from the one for the Metamath Proof Explorer but feel free to copy-paste a citation in either direction in order to cite it.

Label naming conventions

Here are a few of the label naming conventions:

  • Suffixes. We follow the conventions of the Metamath Proof Explorer with a few additions. A biconditional in set.mm which is an implication in iset.mm should have a "r" (for the reverse direction), or "i"/"im" (for the forward direction) appended. A theorem in set.mm which has a decidability condition added should add "dc" to the theorem name. A theorem in set.mm where "nonempty class" is changed to "inhabited class" should add "m" (for member) to the theorem name.

The following table shows some commonly-used abbreviations in labels which are not found in the Metamath Proof Explorer, in alphabetical order. For each abbreviation we provide a mnenomic to help you remember it, the source theorem/assumption defining it, an expression showing what it looks like, whether or not it is a "syntax fragment" (an abbreviation that indicates a particular kind of syntax), and hyperlinks to label examples that use the abbreviation. The abbreviation is bolded if there is a df-NAME definition but the label fragment is not NAME.

AbbreviationMnenomicSource ExpressionSyntax?Example(s)
apapart df-ap 7682 Yes apadd1 7708, apne 7723

  • Community. The Metamath mailing list also covers the Intuitionistic Logic Explorer and is at: https://groups.google.com/forum/#!forum/metamath.
  • (Contributed by Jim Kingdon, 24-Feb-2020.)

    𝜑       𝜑
     
    5.1.2  Definitional examples
     
    Theoremex-or 10560 Example for ax-io 662. Example by David A. Wheeler. (Contributed by Mario Carneiro, 9-May-2015.)
    (2 = 3 ∨ 4 = 4)
     
    Theoremex-an 10561 Example for ax-ia1 104. Example by David A. Wheeler. (Contributed by Mario Carneiro, 9-May-2015.)
    (2 = 2 ∧ 3 = 3)
     
    Theorem1kp2ke3k 10562 Example for df-dec 8478, 1000 + 2000 = 3000.

    This proof disproves (by counterexample) the assertion of Hao Wang, who stated, "There is a theorem in the primitive notation of set theory that corresponds to the arithmetic theorem 1000 + 2000 = 3000. The formula would be forbiddingly long... even if (one) knows the definitions and is asked to simplify the long formula according to them, chances are he will make errors and arrive at some incorrect result." (Hao Wang, "Theory and practice in mathematics" , In Thomas Tymoczko, editor, New Directions in the Philosophy of Mathematics, pp 129-152, Birkauser Boston, Inc., Boston, 1986. (QA8.6.N48). The quote itself is on page 140.)

    This is noted in Metamath: A Computer Language for Pure Mathematics by Norman Megill (2007) section 1.1.3. Megill then states, "A number of writers have conveyed the impression that the kind of absolute rigor provided by Metamath is an impossible dream, suggesting that a complete, formal verification of a typical theorem would take millions of steps in untold volumes of books... These writers assume, however, that in order to achieve the kind of complete formal verification they desire one must break down a proof into individual primitive steps that make direct reference to the axioms. This is not necessary. There is no reason not to make use of previously proved theorems rather than proving them over and over... A hierarchy of theorems and definitions permits an exponential growth in the formula sizes and primitive proof steps to be described with only a linear growth in the number of symbols used. Of course, this is how ordinary informal mathematics is normally done anyway, but with Metamath it can be done with absolute rigor and precision."

    The proof here starts with (2 + 1) = 3, commutes it, and repeatedly multiplies both sides by ten. This is certainly longer than traditional mathematical proofs, e.g., there are a number of steps explicitly shown here to show that we're allowed to do operations such as multiplication. However, while longer, the proof is clearly a manageable size - even though every step is rigorously derived all the way back to the primitive notions of set theory and logic. And while there's a risk of making errors, the many independent verifiers make it much less likely that an incorrect result will be accepted.

    This proof heavily relies on the decimal constructor df-dec 8478 developed by Mario Carneiro in 2015. The underlying Metamath language has an intentionally very small set of primitives; it doesn't even have a built-in construct for numbers. Instead, the digits are defined using these primitives, and the decimal constructor is used to make it easy to express larger numbers as combinations of digits.

    (Contributed by David A. Wheeler, 29-Jun-2016.) (Shortened by Mario Carneiro using the arithmetic algorithm in mmj2, 30-Jun-2016.)

    (1000 + 2000) = 3000
     
    Theoremex-fl 10563 Example for df-fl 9274. Example by David A. Wheeler. (Contributed by Mario Carneiro, 18-Jun-2015.)
    ((⌊‘(3 / 2)) = 1 ∧ (⌊‘-(3 / 2)) = -2)
     
    Theoremex-ceil 10564 Example for df-ceil 9275. (Contributed by AV, 4-Sep-2021.)
    ((⌈‘(3 / 2)) = 2 ∧ (⌈‘-(3 / 2)) = -1)
     
    Theoremex-fac 10565 Example for df-fac 9653. (Contributed by AV, 4-Sep-2021.)
    (!‘5) = 120
     
    Theoremex-bc 10566 Example for df-bc 9675. (Contributed by AV, 4-Sep-2021.)
    (5C3) = 10
     
    Theoremex-dvds 10567 Example for df-dvds 10196: 3 divides into 6. (Contributed by David A. Wheeler, 19-May-2015.)
    3 ∥ 6
     
    Theoremex-gcd 10568 Example for df-gcd 10339. (Contributed by AV, 5-Sep-2021.)
    (-6 gcd 9) = 3
     
    PART 6  SUPPLEMENTARY MATERIAL (USER'S MATHBOXES)
     
    6.1  Mathboxes for user contributions
     
    6.1.1  Mathbox guidelines
     
    Theoremmathbox 10569 (This theorem is a dummy placeholder for these guidelines. The name of this theorem, "mathbox", is hard-coded into the Metamath program to identify the start of the mathbox section for web page generation.)

    A "mathbox" is a user-contributed section that is maintained by its contributor independently from the main part of set.mm.

    For contributors:

    By making a contribution, you agree to release it into the public domain, according to the statement at the beginning of set.mm.

    Mathboxes are provided to help keep your work synchronized with changes in set.mm, but they shouldn't be depended on as a permanent archive. If you want to preserve your original contribution, it is your responsibility to keep your own copy of it along with the version of set.mm that works with it.

    Guidelines:

    1. If at all possible, please use only nullary class constants for new definitions.

    2. Try to follow the style of the rest of set.mm. Each $p and $a statement must be immediately preceded with the comment that will be shown on its web page description. The metamath program command "write source set.mm /rewrap" will take care of wrapping comment lines and indentation conventions. All mathbox content will be on public display and should hopefully reflect the overall quality of the website.

    3. Before submitting a revised mathbox, please make sure it verifies against the current set.mm.

    4. Mathboxes should be independent i.e. the proofs should verify with all other mathboxes removed. If you need a theorem from another mathbox, that is fine (and encouraged), but let me know, so I can move the theorem to the main section. One way avoid undesired accidental use of other mathbox theorems is to develop your mathbox using a modified set.mm that has mathboxes removed.

    Notes:

    1. We may decide to move some theorems to the main part of set.mm for general use.

    2. We may make changes to mathboxes to maintain the overall quality of set.mm. Normally we will let you know if a change might impact what you are working on.

    3. If you use theorems from another user's mathbox, we don't provide assurance that they are based on correct or consistent $a statements. (If you find such a problem, please let us know so it can be corrected.) (Contributed by NM, 20-Feb-2007.) (New usage is discouraged.)

    𝜑       𝜑
     
    6.2  Mathbox for BJ
     
    6.2.1  Propositional calculus
     
    Theoremnnexmid 10570 Double negation of excluded middle. Intuitionistic logic refutes the negation of excluded middle (but, of course, does not prove excluded middle) for any formula. (Contributed by BJ, 9-Oct-2019.)
    ¬ ¬ (𝜑 ∨ ¬ 𝜑)
     
    Theoremnndc 10571 Double negation of decidability of a formula. Intuitionistic logic refutes undecidability (but, of course, does not prove decidability) of any formula. (Contributed by BJ, 9-Oct-2019.)
    ¬ ¬ DECID 𝜑
     
    Theoremdcdc 10572 Decidability of a proposition is decidable if and only if that proposition is decidable. DECID is idempotent. (Contributed by BJ, 9-Oct-2019.)
    (DECID DECID 𝜑DECID 𝜑)
     
    6.2.2  Predicate calculus
     
    Theorembj-ex 10573* Existential generalization. (Contributed by BJ, 8-Dec-2019.) Proof modification is discouraged because there are shorter proofs, but using less basic results (like exlimiv 1529 and 19.9ht 1572 or 19.23ht 1426). (Proof modification is discouraged.)
    (∃𝑥𝜑𝜑)
     
    Theorembj-hbalt 10574 Closed form of hbal 1406 (copied from set.mm). (Contributed by BJ, 2-May-2019.)
    (∀𝑦(𝜑 → ∀𝑥𝜑) → (∀𝑦𝜑 → ∀𝑥𝑦𝜑))
     
    Theorembj-nfalt 10575 Closed form of nfal 1508 (copied from set.mm). (Contributed by BJ, 2-May-2019.)
    (∀𝑥𝑦𝜑 → Ⅎ𝑦𝑥𝜑)
     
    Theoremspimd 10576 Deduction form of spim 1666. (Contributed by BJ, 17-Oct-2019.)
    (𝜑 → Ⅎ𝑥𝜒)    &   (𝜑 → ∀𝑥(𝑥 = 𝑦 → (𝜓𝜒)))       (𝜑 → (∀𝑥𝜓𝜒))
     
    Theorem2spim 10577* Double substitution, as in spim 1666. (Contributed by BJ, 17-Oct-2019.)
    𝑥𝜒    &   𝑧𝜒    &   ((𝑥 = 𝑦𝑧 = 𝑡) → (𝜓𝜒))       (∀𝑧𝑥𝜓𝜒)
     
    Theoremch2var 10578* Implicit substitution of 𝑦 for 𝑥 and 𝑡 for 𝑧 into a theorem. (Contributed by BJ, 17-Oct-2019.)
    𝑥𝜓    &   𝑧𝜓    &   ((𝑥 = 𝑦𝑧 = 𝑡) → (𝜑𝜓))    &   𝜑       𝜓
     
    Theoremch2varv 10579* Version of ch2var 10578 with non-freeness hypotheses replaced by DV conditions. (Contributed by BJ, 17-Oct-2019.)
    ((𝑥 = 𝑦𝑧 = 𝑡) → (𝜑𝜓))    &   𝜑       𝜓
     
    Theorembj-exlimmp 10580 Lemma for bj-vtoclgf 10586. (Contributed by BJ, 21-Nov-2019.) (Proof modification is discouraged.)
    𝑥𝜓    &   (𝜒𝜑)       (∀𝑥(𝜒 → (𝜑𝜓)) → (∃𝑥𝜒𝜓))
     
    Theorembj-exlimmpi 10581 Lemma for bj-vtoclgf 10586. (Contributed by BJ, 21-Nov-2019.) (Proof modification is discouraged.)
    𝑥𝜓    &   (𝜒𝜑)    &   (𝜒 → (𝜑𝜓))       (∃𝑥𝜒𝜓)
     
    Theorembj-sbimedh 10582 A strengthening of sbiedh 1710 (same proof). (Contributed by BJ, 16-Dec-2019.)
    (𝜑 → ∀𝑥𝜑)    &   (𝜑 → (𝜒 → ∀𝑥𝜒))    &   (𝜑 → (𝑥 = 𝑦 → (𝜓𝜒)))       (𝜑 → ([𝑦 / 𝑥]𝜓𝜒))
     
    Theorembj-sbimeh 10583 A strengthening of sbieh 1713 (same proof). (Contributed by BJ, 16-Dec-2019.)
    (𝜓 → ∀𝑥𝜓)    &   (𝑥 = 𝑦 → (𝜑𝜓))       ([𝑦 / 𝑥]𝜑𝜓)
     
    Theorembj-sbime 10584 A strengthening of sbie 1714 (same proof). (Contributed by BJ, 16-Dec-2019.)
    𝑥𝜓    &   (𝑥 = 𝑦 → (𝜑𝜓))       ([𝑦 / 𝑥]𝜑𝜓)
     
    6.2.3  Extensionality

    Various utility theorems using FOL and extensionality.

     
    Theorembj-vtoclgft 10585 Weakening two hypotheses of vtoclgf 2657. (Contributed by BJ, 21-Nov-2019.)
    𝑥𝐴    &   𝑥𝜓    &   (𝑥 = 𝐴𝜑)       (∀𝑥(𝑥 = 𝐴 → (𝜑𝜓)) → (𝐴𝑉𝜓))
     
    Theorembj-vtoclgf 10586 Weakening two hypotheses of vtoclgf 2657. (Contributed by BJ, 21-Nov-2019.)
    𝑥𝐴    &   𝑥𝜓    &   (𝑥 = 𝐴𝜑)    &   (𝑥 = 𝐴 → (𝜑𝜓))       (𝐴𝑉𝜓)
     
    Theoremelabgf0 10587 Lemma for elabgf 2736. (Contributed by BJ, 21-Nov-2019.)
    (𝑥 = 𝐴 → (𝐴 ∈ {𝑥𝜑} ↔ 𝜑))
     
    Theoremelabgft1 10588 One implication of elabgf 2736, in closed form. (Contributed by BJ, 21-Nov-2019.)
    𝑥𝐴    &   𝑥𝜓       (∀𝑥(𝑥 = 𝐴 → (𝜑𝜓)) → (𝐴 ∈ {𝑥𝜑} → 𝜓))
     
    Theoremelabgf1 10589 One implication of elabgf 2736. (Contributed by BJ, 21-Nov-2019.)
    𝑥𝐴    &   𝑥𝜓    &   (𝑥 = 𝐴 → (𝜑𝜓))       (𝐴 ∈ {𝑥𝜑} → 𝜓)
     
    Theoremelabgf2 10590 One implication of elabgf 2736. (Contributed by BJ, 21-Nov-2019.)
    𝑥𝐴    &   𝑥𝜓    &   (𝑥 = 𝐴 → (𝜓𝜑))       (𝐴𝐵 → (𝜓𝐴 ∈ {𝑥𝜑}))
     
    Theoremelabf1 10591* One implication of elabf 2737. (Contributed by BJ, 21-Nov-2019.)
    𝑥𝜓    &   (𝑥 = 𝐴 → (𝜑𝜓))       (𝐴 ∈ {𝑥𝜑} → 𝜓)
     
    Theoremelabf2 10592* One implication of elabf 2737. (Contributed by BJ, 21-Nov-2019.)
    𝑥𝜓    &   𝐴 ∈ V    &   (𝑥 = 𝐴 → (𝜓𝜑))       (𝜓𝐴 ∈ {𝑥𝜑})
     
    Theoremelab1 10593* One implication of elab 2738. (Contributed by BJ, 21-Nov-2019.)
    (𝑥 = 𝐴 → (𝜑𝜓))       (𝐴 ∈ {𝑥𝜑} → 𝜓)
     
    Theoremelab2a 10594* One implication of elab 2738. (Contributed by BJ, 21-Nov-2019.)
    𝐴 ∈ V    &   (𝑥 = 𝐴 → (𝜓𝜑))       (𝜓𝐴 ∈ {𝑥𝜑})
     
    Theoremelabg2 10595* One implication of elabg 2739. (Contributed by BJ, 21-Nov-2019.)
    (𝑥 = 𝐴 → (𝜓𝜑))       (𝐴𝑉 → (𝜓𝐴 ∈ {𝑥𝜑}))
     
    Theorembj-rspgt 10596 Restricted specialization, generalized. Weakens a hypothesis of rspccv 2698 and seems to have a shorter proof. (Contributed by BJ, 21-Nov-2019.)
    𝑥𝐴    &   𝑥𝐵    &   𝑥𝜓       (∀𝑥(𝑥 = 𝐴 → (𝜑𝜓)) → (∀𝑥𝐵 𝜑 → (𝐴𝐵𝜓)))
     
    Theorembj-rspg 10597 Restricted specialization, generalized. Weakens a hypothesis of rspccv 2698 and seems to have a shorter proof. (Contributed by BJ, 21-Nov-2019.)
    𝑥𝐴    &   𝑥𝐵    &   𝑥𝜓    &   (𝑥 = 𝐴 → (𝜑𝜓))       (∀𝑥𝐵 𝜑 → (𝐴𝐵𝜓))
     
    Theoremcbvrald 10598* Rule used to change bound variables, using implicit substitution. (Contributed by BJ, 22-Nov-2019.)
    𝑥𝜑    &   𝑦𝜑    &   (𝜑 → Ⅎ𝑦𝜓)    &   (𝜑 → Ⅎ𝑥𝜒)    &   (𝜑 → (𝑥 = 𝑦 → (𝜓𝜒)))       (𝜑 → (∀𝑥𝐴 𝜓 ↔ ∀𝑦𝐴 𝜒))
     
    Theorembj-intabssel 10599 Version of intss1 3651 using a class abstraction and explicit substitution. (Contributed by BJ, 29-Nov-2019.)
    𝑥𝐴       (𝐴𝑉 → ([𝐴 / 𝑥]𝜑 {𝑥𝜑} ⊆ 𝐴))
     
    Theorembj-intabssel1 10600 Version of intss1 3651 using a class abstraction and implicit substitution. Closed form of intmin3 3663. (Contributed by BJ, 29-Nov-2019.)
    𝑥𝐴    &   𝑥𝜓    &   (𝑥 = 𝐴 → (𝜓𝜑))       (𝐴𝑉 → (𝜓 {𝑥𝜑} ⊆ 𝐴))
        < Previous  Next >

    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  Next >