HomeHome Intuitionistic Logic Explorer
Theorem List (p. 44 of 108)
< Previous  Next >
Browser slow? Try the
Unicode version.

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

Theorem List for Intuitionistic Logic Explorer - 4301-4400   *Has distinct variable group(s)
TypeLabelDescription
Statement
 
Theoremsuc11 4301 The successor operation behaves like a one-to-one function. Compare Exercise 16 of [Enderton] p. 194. (Contributed by NM, 3-Sep-2003.)
 |-  ( ( A  e.  On  /\  B  e.  On )  ->  ( suc  A  =  suc  B  <->  A  =  B ) )
 
Theoremdtruex 4302* At least two sets exist (or in terms of first-order logic, the universe of discourse has two or more objects). Although dtruarb 3962 can also be summarized as "at least two sets exist", the difference is that dtruarb 3962 shows the existence of two sets which are not equal to each other, but this theorem says that given a specific  y, we can construct a set  x which does not equal it. (Contributed by Jim Kingdon, 29-Dec-2018.)
 |- 
 E. x  -.  x  =  y
 
Theoremdtru 4303* At least two sets exist (or in terms of first-order logic, the universe of discourse has two or more objects). If we assumed the law of the excluded middle this would be equivalent to dtruex 4302. (Contributed by Jim Kingdon, 29-Dec-2018.)
 |- 
 -.  A. x  x  =  y
 
Theoremeunex 4304 Existential uniqueness implies there is a value for which the wff argument is false. (Contributed by Jim Kingdon, 29-Dec-2018.)
 |-  ( E! x ph  ->  E. x  -.  ph )
 
Theoremordsoexmid 4305 Weak linearity of ordinals implies the law of the excluded middle (that is, decidability of an arbitrary proposition). (Contributed by Mario Carneiro and Jim Kingdon, 29-Jan-2019.)
 |- 
 _E  Or  On   =>    |-  ( ph  \/  -.  ph )
 
Theoremordsuc 4306 The successor of an ordinal class is ordinal. (Contributed by NM, 3-Apr-1995.) (Constructive proof by Mario Carneiro and Jim Kingdon, 20-Jul-2019.)
 |-  ( Ord  A  <->  Ord  suc  A )
 
Theoremonsucuni2 4307 A successor ordinal is the successor of its union. (Contributed by NM, 10-Dec-2004.) (Proof shortened by Andrew Salmon, 27-Aug-2011.)
 |-  ( ( A  e.  On  /\  A  =  suc  B )  ->  suc  U. A  =  A )
 
Theorem0elsucexmid 4308* If the successor of any ordinal class contains the empty set, excluded middle follows. (Contributed by Jim Kingdon, 3-Sep-2021.)
 |- 
 A. x  e.  On  (/) 
 e.  suc  x   =>    |-  ( ph  \/  -.  ph )
 
Theoremnlimsucg 4309 A successor is not a limit ordinal. (Contributed by NM, 25-Mar-1995.) (Proof shortened by Andrew Salmon, 27-Aug-2011.)
 |-  ( A  e.  V  ->  -.  Lim  suc  A )
 
Theoremordpwsucss 4310 The collection of ordinals in the power class of an ordinal is a superset of its successor.

We can think of  ( ~P A  i^i  On ) as another possible definition of successor, which would be equivalent to df-suc 4126 given excluded middle. It is an ordinal, and has some successor-like properties. For example, if  A  e.  On then both  U. suc  A  =  A (onunisuci 4187) and  U. { x  e.  On  |  x  C_  A }  =  A (onuniss2 4256).

Constructively  ( ~P A  i^i  On ) and  suc  A cannot be shown to be equivalent (as proved at ordpwsucexmid 4313). (Contributed by Jim Kingdon, 21-Jul-2019.)

 |-  ( Ord  A  ->  suc 
 A  C_  ( ~P A  i^i  On ) )
 
Theoremonnmin 4311 No member of a set of ordinal numbers belongs to its minimum. (Contributed by NM, 2-Feb-1997.) (Constructive proof by Mario Carneiro and Jim Kingdon, 21-Jul-2019.)
 |-  ( ( A  C_  On  /\  B  e.  A )  ->  -.  B  e.  |^|
 A )
 
Theoremssnel 4312 Relationship between subset and elementhood. In the context of ordinals this can be seen as an ordering law. (Contributed by Jim Kingdon, 22-Jul-2019.)
 |-  ( A  C_  B  ->  -.  B  e.  A )
 
Theoremordpwsucexmid 4313* The subset in ordpwsucss 4310 cannot be equality. That is, strengthening it to equality implies excluded middle. (Contributed by Jim Kingdon, 30-Jul-2019.)
 |- 
 A. x  e.  On  suc 
 x  =  ( ~P x  i^i  On )   =>    |-  ( ph  \/  -.  ph )
 
Theoremordtri2or2exmid 4314* Ordinal trichotomy implies excluded middle. (Contributed by Jim Kingdon, 29-Aug-2021.)
 |- 
 A. x  e.  On  A. y  e.  On  ( x  C_  y  \/  y  C_  x )   =>    |-  ( ph  \/  -.  ph )
 
Theoremonintexmid 4315* If the intersection (infimum) of an inhabited class of ordinal numbers belongs to the class, excluded middle follows. The hypothesis would be provable given excluded middle. (Contributed by Mario Carneiro and Jim Kingdon, 29-Aug-2021.)
 |-  ( ( y  C_  On  /\  E. x  x  e.  y )  ->  |^| y  e.  y
 )   =>    |-  ( ph  \/  -.  ph )
 
Theoremzfregfr 4316 The epsilon relation is well-founded on any class. (Contributed by NM, 26-Nov-1995.)
 |- 
 _E  Fr  A
 
Theoremordfr 4317 Epsilon is well-founded on an ordinal class. (Contributed by NM, 22-Apr-1994.)
 |-  ( Ord  A  ->  _E 
 Fr  A )
 
Theoremordwe 4318 Epsilon well-orders every ordinal. Proposition 7.4 of [TakeutiZaring] p. 36. (Contributed by NM, 3-Apr-1994.)
 |-  ( Ord  A  ->  _E 
 We  A )
 
Theoremwetriext 4319* A trichotomous well-order is extensional. (Contributed by Jim Kingdon, 26-Sep-2021.)
 |-  ( ph  ->  R  We  A )   &    |-  ( ph  ->  A  e.  V )   &    |-  ( ph  ->  A. a  e.  A  A. b  e.  A  ( a R b  \/  a  =  b  \/  b R a ) )   &    |-  ( ph  ->  B  e.  A )   &    |-  ( ph  ->  C  e.  A )   &    |-  ( ph  ->  A. z  e.  A  ( z R B  <->  z R C ) )   =>    |-  ( ph  ->  B  =  C )
 
Theoremwessep 4320 A subset of a set well-ordered by set membership is well-ordered by set membership. (Contributed by Jim Kingdon, 30-Sep-2021.)
 |-  ( (  _E  We  A  /\  B  C_  A )  ->  _E  We  B )
 
Theoremreg3exmidlemwe 4321* Lemma for reg3exmid 4322. Our counterexample  A satisfies  We. (Contributed by Jim Kingdon, 3-Oct-2021.)
 |-  A  =  { x  e.  { (/) ,  { (/) } }  |  ( x  =  { (/)
 }  \/  ( x  =  (/)  /\  ph ) ) }   =>    |- 
 _E  We  A
 
Theoremreg3exmid 4322* If any inhabited set satisfying df-wetr 4089 for  _E has a minimal element, excluded middle follows. (Contributed by Jim Kingdon, 3-Oct-2021.)
 |-  ( (  _E  We  z  /\  E. w  w  e.  z )  ->  E. x  e.  z  A. y  e.  z  x  C_  y )   =>    |-  ( ph  \/  -.  ph )
 
2.5.3  Transfinite induction
 
Theoremtfi 4323* The Principle of Transfinite Induction. Theorem 7.17 of [TakeutiZaring] p. 39. This principle states that if  A is a class of ordinal numbers with the property that every ordinal number included in  A also belongs to  A, then every ordinal number is in  A.

(Contributed by NM, 18-Feb-2004.)

 |-  ( ( A  C_  On  /\  A. x  e. 
 On  ( x  C_  A  ->  x  e.  A ) )  ->  A  =  On )
 
Theoremtfis 4324* Transfinite Induction Schema. If all ordinal numbers less than a given number  x have a property (induction hypothesis), then all ordinal numbers have the property (conclusion). Exercise 25 of [Enderton] p. 200. (Contributed by NM, 1-Aug-1994.) (Revised by Mario Carneiro, 20-Nov-2016.)
 |-  ( x  e.  On  ->  ( A. y  e.  x  [ y  /  x ] ph  ->  ph )
 )   =>    |-  ( x  e.  On  -> 
 ph )
 
Theoremtfis2f 4325* Transfinite Induction Schema, using implicit substitution. (Contributed by NM, 18-Aug-1994.)
 |- 
 F/ x ps   &    |-  ( x  =  y  ->  (
 ph 
 <->  ps ) )   &    |-  ( x  e.  On  ->  (
 A. y  e.  x  ps  ->  ph ) )   =>    |-  ( x  e. 
 On  ->  ph )
 
Theoremtfis2 4326* Transfinite Induction Schema, using implicit substitution. (Contributed by NM, 18-Aug-1994.)
 |-  ( x  =  y 
 ->  ( ph  <->  ps ) )   &    |-  ( x  e.  On  ->  (
 A. y  e.  x  ps  ->  ph ) )   =>    |-  ( x  e. 
 On  ->  ph )
 
Theoremtfis3 4327* Transfinite Induction Schema, using implicit substitution. (Contributed by NM, 4-Nov-2003.)
 |-  ( x  =  y 
 ->  ( ph  <->  ps ) )   &    |-  ( x  =  A  ->  (
 ph 
 <->  ch ) )   &    |-  ( x  e.  On  ->  (
 A. y  e.  x  ps  ->  ph ) )   =>    |-  ( A  e.  On  ->  ch )
 
Theoremtfisi 4328* A transfinite induction scheme in "implicit" form where the induction is done on an object derived from the object of interest. (Contributed by Stefan O'Rear, 24-Aug-2015.)
 |-  ( ph  ->  A  e.  V )   &    |-  ( ph  ->  T  e.  On )   &    |-  (
 ( ph  /\  ( R  e.  On  /\  R  C_  T )  /\  A. y ( S  e.  R  ->  ch ) )  ->  ps )   &    |-  ( x  =  y  ->  ( ps  <->  ch ) )   &    |-  ( x  =  A  ->  ( ps  <->  th ) )   &    |-  ( x  =  y  ->  R  =  S )   &    |-  ( x  =  A  ->  R  =  T )   =>    |-  ( ph  ->  th )
 
2.6  IZF Set Theory - add the Axiom of Infinity
 
2.6.1  Introduce the Axiom of Infinity
 
Axiomax-iinf 4329* Axiom of Infinity. Axiom 5 of [Crosilla] p. "Axioms of CZF and IZF". (Contributed by Jim Kingdon, 16-Nov-2018.)
 |- 
 E. x ( (/)  e.  x  /\  A. y
 ( y  e.  x  ->  suc  y  e.  x ) )
 
Theoremzfinf2 4330* A standard version of the Axiom of Infinity, using definitions to abbreviate. Axiom Inf of [BellMachover] p. 472. (Contributed by NM, 30-Aug-1993.)
 |- 
 E. x ( (/)  e.  x  /\  A. y  e.  x  suc  y  e.  x )
 
2.6.2  The natural numbers (i.e. finite ordinals)
 
Syntaxcom 4331 Extend class notation to include the class of natural numbers.
 class  om
 
Definitiondf-iom 4332* Define the class of natural numbers as the smallest inductive set, which is valid provided we assume the Axiom of Infinity. Definition 6.3 of [Eisenberg] p. 82.

Note: the natural numbers  om are a subset of the ordinal numbers df-on 4123. Later, when we define complex numbers, we will be able to also define a subset of the complex numbers with analogous properties and operations, but they will be different sets. (Contributed by NM, 6-Aug-1994.) Use its alias dfom3 4333 instead for naming consistency with set.mm. (New usage is discouraged.)

 |- 
 om  =  |^| { x  |  ( (/)  e.  x  /\  A. y  e.  x  suc  y  e.  x ) }
 
Theoremdfom3 4333* Alias for df-iom 4332. Use it instead of df-iom 4332 for naming consistency with set.mm. (Contributed by NM, 6-Aug-1994.)
 |- 
 om  =  |^| { x  |  ( (/)  e.  x  /\  A. y  e.  x  suc  y  e.  x ) }
 
Theoremomex 4334 The existence of omega (the class of natural numbers). Axiom 7 of [TakeutiZaring] p. 43. (Contributed by NM, 6-Aug-1994.)
 |- 
 om  e.  _V
 
2.6.3  Peano's postulates
 
Theorempeano1 4335 Zero is a natural number. One of Peano's five postulates for arithmetic. Proposition 7.30(1) of [TakeutiZaring] p. 42. (Contributed by NM, 15-May-1994.)
 |-  (/)  e.  om
 
Theorempeano2 4336 The successor of any natural number is a natural number. One of Peano's five postulates for arithmetic. Proposition 7.30(2) of [TakeutiZaring] p. 42. (Contributed by NM, 3-Sep-2003.)
 |-  ( A  e.  om  ->  suc  A  e.  om )
 
Theorempeano3 4337 The successor of any natural number is not zero. One of Peano's five postulates for arithmetic. Proposition 7.30(3) of [TakeutiZaring] p. 42. (Contributed by NM, 3-Sep-2003.)
 |-  ( A  e.  om  ->  suc  A  =/=  (/) )
 
Theorempeano4 4338 Two natural numbers are equal iff their successors are equal, i.e. the successor function is one-to-one. One of Peano's five postulates for arithmetic. Proposition 7.30(4) of [TakeutiZaring] p. 43. (Contributed by NM, 3-Sep-2003.)
 |-  ( ( A  e.  om 
 /\  B  e.  om )  ->  ( suc  A  =  suc  B  <->  A  =  B ) )
 
Theorempeano5 4339* The induction postulate: any class containing zero and closed under the successor operation contains all natural numbers. One of Peano's five postulates for arithmetic. Proposition 7.30(5) of [TakeutiZaring] p. 43. The more traditional statement of mathematical induction as a theorem schema, with a basis and an induction step, is derived from this theorem as theorem findes 4344. (Contributed by NM, 18-Feb-2004.)
 |-  ( ( (/)  e.  A  /\  A. x  e.  om  ( x  e.  A  ->  suc  x  e.  A ) )  ->  om  C_  A )
 
2.6.4  Finite induction (for finite ordinals)
 
Theoremfind 4340* The Principle of Finite Induction (mathematical induction). Corollary 7.31 of [TakeutiZaring] p. 43. The simpler hypothesis shown here was suggested in an email from "Colin" on 1-Oct-2001. The hypothesis states that  A is a set of natural numbers, zero belongs to 
A, and given any member of  A the member's successor also belongs to  A. The conclusion is that every natural number is in  A. (Contributed by NM, 22-Feb-2004.) (Proof shortened by Andrew Salmon, 27-Aug-2011.)
 |-  ( A  C_  om  /\  (/) 
 e.  A  /\  A. x  e.  A  suc  x  e.  A )   =>    |-  A  =  om
 
Theoremfinds 4341* Principle of Finite Induction (inference schema), using implicit substitutions. The first four hypotheses establish the substitutions we need. The last two are the basis and the induction step. Theorem Schema 22 of [Suppes] p. 136. This is Metamath 100 proof #74. (Contributed by NM, 14-Apr-1995.)
 |-  ( x  =  (/)  ->  ( ph  <->  ps ) )   &    |-  ( x  =  y  ->  (
 ph 
 <->  ch ) )   &    |-  ( x  =  suc  y  ->  ( ph  <->  th ) )   &    |-  ( x  =  A  ->  (
 ph 
 <->  ta ) )   &    |-  ps   &    |-  (
 y  e.  om  ->  ( ch  ->  th )
 )   =>    |-  ( A  e.  om  ->  ta )
 
Theoremfinds2 4342* Principle of Finite Induction (inference schema), using implicit substitutions. The first three hypotheses establish the substitutions we need. The last two are the basis and the induction step. Theorem Schema 22 of [Suppes] p. 136. (Contributed by NM, 29-Nov-2002.)
 |-  ( x  =  (/)  ->  ( ph  <->  ps ) )   &    |-  ( x  =  y  ->  (
 ph 
 <->  ch ) )   &    |-  ( x  =  suc  y  ->  ( ph  <->  th ) )   &    |-  ( ta  ->  ps )   &    |-  ( y  e. 
 om  ->  ( ta  ->  ( ch  ->  th )
 ) )   =>    |-  ( x  e.  om  ->  ( ta  ->  ph )
 )
 
Theoremfinds1 4343* Principle of Finite Induction (inference schema), using implicit substitutions. The first three hypotheses establish the substitutions we need. The last two are the basis and the induction step. Theorem Schema 22 of [Suppes] p. 136. (Contributed by NM, 22-Mar-2006.)
 |-  ( x  =  (/)  ->  ( ph  <->  ps ) )   &    |-  ( x  =  y  ->  (
 ph 
 <->  ch ) )   &    |-  ( x  =  suc  y  ->  ( ph  <->  th ) )   &    |-  ps   &    |-  (
 y  e.  om  ->  ( ch  ->  th )
 )   =>    |-  ( x  e.  om  -> 
 ph )
 
Theoremfindes 4344 Finite induction with explicit substitution. The first hypothesis is the basis and the second is the induction step. Theorem Schema 22 of [Suppes] p. 136. This is an alternative for Metamath 100 proof #74. (Contributed by Raph Levien, 9-Jul-2003.)
 |-  [. (/)  /  x ]. ph   &    |-  ( x  e.  om  ->  (
 ph  ->  [. suc  x  /  x ]. ph ) )   =>    |-  ( x  e.  om  ->  ph )
 
2.6.5  The Natural Numbers (continued)
 
Theoremnn0suc 4345* A natural number is either 0 or a successor. Similar theorems for arbitrary sets or real numbers will not be provable (without the law of the excluded middle), but equality of natural numbers is decidable. (Contributed by NM, 27-May-1998.)
 |-  ( A  e.  om  ->  ( A  =  (/)  \/ 
 E. x  e.  om  A  =  suc  x ) )
 
Theoremelnn 4346 A member of a natural number is a natural number. (Contributed by NM, 21-Jun-1998.)
 |-  ( ( A  e.  B  /\  B  e.  om )  ->  A  e.  om )
 
Theoremordom 4347 Omega is ordinal. Theorem 7.32 of [TakeutiZaring] p. 43. (Contributed by NM, 18-Oct-1995.)
 |- 
 Ord  om
 
Theoremomelon2 4348 Omega is an ordinal number. (Contributed by Mario Carneiro, 30-Jan-2013.)
 |-  ( om  e.  _V  ->  om  e.  On )
 
Theoremomelon 4349 Omega is an ordinal number. (Contributed by NM, 10-May-1998.) (Revised by Mario Carneiro, 30-Jan-2013.)
 |- 
 om  e.  On
 
Theoremnnon 4350 A natural number is an ordinal number. (Contributed by NM, 27-Jun-1994.)
 |-  ( A  e.  om  ->  A  e.  On )
 
Theoremnnoni 4351 A natural number is an ordinal number. (Contributed by NM, 27-Jun-1994.)
 |-  A  e.  om   =>    |-  A  e.  On
 
Theoremnnord 4352 A natural number is ordinal. (Contributed by NM, 17-Oct-1995.)
 |-  ( A  e.  om  ->  Ord  A )
 
Theoremomsson 4353 Omega is a subset of  On. (Contributed by NM, 13-Jun-1994.)
 |- 
 om  C_  On
 
Theoremlimom 4354 Omega is a limit ordinal. Theorem 2.8 of [BellMachover] p. 473. (Contributed by NM, 26-Mar-1995.) (Proof rewritten by Jim Kingdon, 5-Jan-2019.)
 |- 
 Lim  om
 
Theorempeano2b 4355 A class belongs to omega iff its successor does. (Contributed by NM, 3-Dec-1995.)
 |-  ( A  e.  om  <->  suc  A  e.  om )
 
Theoremnnsuc 4356* A nonzero natural number is a successor. (Contributed by NM, 18-Feb-2004.)
 |-  ( ( A  e.  om 
 /\  A  =/=  (/) )  ->  E. x  e.  om  A  =  suc  x )
 
Theoremnndceq0 4357 A natural number is either zero or nonzero. Decidable equality for natural numbers is a special case of the law of the excluded middle which holds in most constructive set theories including ours. (Contributed by Jim Kingdon, 5-Jan-2019.)
 |-  ( A  e.  om  -> DECID  A  =  (/) )
 
Theorem0elnn 4358 A natural number is either the empty set or has the empty set as an element. (Contributed by Jim Kingdon, 23-Aug-2019.)
 |-  ( A  e.  om  ->  ( A  =  (/)  \/  (/)  e.  A ) )
 
Theoremnn0eln0 4359 A natural number is nonempty iff it contains the empty set. Although in constructive mathematics it is generally more natural to work with inhabited sets and ignore the whole concept of nonempty sets, in the specific case of natural numbers this theorem may be helpful in converting proofs which were written assuming excluded middle. (Contributed by Jim Kingdon, 28-Aug-2019.)
 |-  ( A  e.  om  ->  ( (/)  e.  A  <->  A  =/=  (/) ) )
 
Theoremnnregexmid 4360* If inhabited sets of natural numbers always have minimal elements, excluded middle follows. The argument is essentially the same as regexmid 4278 and the larger lesson is that although natural numbers may behave "non-constructively" even in a constructive set theory (for example see nndceq 6100 or nntri3or 6095), sets of natural numbers are a different animal. (Contributed by Jim Kingdon, 6-Sep-2019.)
 |-  ( ( x  C_  om 
 /\  E. y  y  e.  x )  ->  E. y
 ( y  e.  x  /\  A. z ( z  e.  y  ->  -.  z  e.  x ) ) )   =>    |-  ( ph  \/  -.  ph )
 
2.6.6  Relations
 
Syntaxcxp 4361 Extend the definition of a class to include the cross product.
 class  ( A  X.  B )
 
Syntaxccnv 4362 Extend the definition of a class to include the converse of a class.
 class  `' A
 
Syntaxcdm 4363 Extend the definition of a class to include the domain of a class.
 class  dom  A
 
Syntaxcrn 4364 Extend the definition of a class to include the range of a class.
 class  ran  A
 
Syntaxcres 4365 Extend the definition of a class to include the restriction of a class. (Read: The restriction of  A to  B.)
 class  ( A  |`  B )
 
Syntaxcima 4366 Extend the definition of a class to include the image of a class. (Read: The image of  B under  A.)
 class  ( A " B )
 
Syntaxccom 4367 Extend the definition of a class to include the composition of two classes. (Read: The composition of  A and  B.)
 class  ( A  o.  B )
 
Syntaxwrel 4368 Extend the definition of a wff to include the relation predicate. (Read:  A is a relation.)
 wff  Rel  A
 
Definitiondf-xp 4369* Define the cross product of two classes. Definition 9.11 of [Quine] p. 64. For example, ( { 1 , 5 }  X. { 2 , 7 } ) = ( {  <. 1 , 2  >.,  <. 1 , 7  >. }  u. {  <. 5 , 2  >.,  <. 5 , 7  >. } ) . Another example is that the set of rational numbers are defined in using the cross-product ( Z  X. N ) ; the left- and right-hand sides of the cross-product represent the top (integer) and bottom (natural) numbers of a fraction. (Contributed by NM, 4-Jul-1994.)
 |-  ( A  X.  B )  =  { <. x ,  y >.  |  ( x  e.  A  /\  y  e.  B ) }
 
Definitiondf-rel 4370 Define the relation predicate. Definition 6.4(1) of [TakeutiZaring] p. 23. For alternate definitions, see dfrel2 4791 and dfrel3 4798. (Contributed by NM, 1-Aug-1994.)
 |-  ( Rel  A  <->  A  C_  ( _V 
 X.  _V ) )
 
Definitiondf-cnv 4371* Define the converse of a class. Definition 9.12 of [Quine] p. 64. The converse of a binary relation swaps its arguments, i.e., if  A  e. 
_V and  B  e.  _V then  ( A `' R B  <-> 
B R A ), as proven in brcnv 4536 (see df-br 3786 and df-rel 4370 for more on relations). For example,  `' {  <. 2 , 6  >.,  <. 3 , 9  >. } = {  <. 6 , 2  >.,  <. 9 , 3  >. } . We use Quine's breve accent (smile) notation. Like Quine, we use it as a prefix, which eliminates the need for parentheses. Many authors use the postfix superscript "to the minus one." "Converse" is Quine's terminology; some authors call it "inverse," especially when the argument is a function. (Contributed by NM, 4-Jul-1994.)
 |-  `' A  =  { <. x ,  y >.  |  y A x }
 
Definitiondf-co 4372* Define the composition of two classes. Definition 6.6(3) of [TakeutiZaring] p. 24. Note that Definition 7 of [Suppes] p. 63 reverses  A and  B, uses a slash instead of  o., and calls the operation "relative product." (Contributed by NM, 4-Jul-1994.)
 |-  ( A  o.  B )  =  { <. x ,  y >.  |  E. z
 ( x B z 
 /\  z A y ) }
 
Definitiondf-dm 4373* Define the domain of a class. Definition 3 of [Suppes] p. 59. For example, F = {  <. 2 , 6  >.,  <. 3 , 9  >. }  -> dom F = { 2 , 3 } . Contrast with range (defined in df-rn 4374). For alternate definitions see dfdm2 4872, dfdm3 4540, and dfdm4 4545. The notation " dom " is used by Enderton; other authors sometimes use script D. (Contributed by NM, 1-Aug-1994.)
 |- 
 dom  A  =  { x  |  E. y  x A y }
 
Definitiondf-rn 4374 Define the range of a class. For example, F = {  <. 2 , 6  >.,  <. 3 , 9  >. } -> ran F = { 6 , 9 } . Contrast with domain (defined in df-dm 4373). For alternate definitions, see dfrn2 4541, dfrn3 4542, and dfrn4 4801. The notation " ran " is used by Enderton; other authors sometimes use script R or script W. (Contributed by NM, 1-Aug-1994.)
 |- 
 ran  A  =  dom  `' A
 
Definitiondf-res 4375 Define the restriction of a class. Definition 6.6(1) of [TakeutiZaring] p. 24. For example ( F = {  <. 2 , 6 
>.,  <. 3 , 9  >. }  /\ B = { 1 , 2 } ) -> ( F  |` B ) = {  <. 2 , 6  >. } . (Contributed by NM, 2-Aug-1994.)
 |-  ( A  |`  B )  =  ( A  i^i  ( B  X.  _V )
 )
 
Definitiondf-ima 4376 Define the image of a class (as restricted by another class). Definition 6.6(2) of [TakeutiZaring] p. 24. For example, ( F = {  <. 2 , 6  >.,  <. 3 , 9  >. } /\ B = { 1 , 2 } ) -> ( F  " B ) = { 6 } . Contrast with restriction (df-res 4375) and range (df-rn 4374). For an alternate definition, see dfima2 4690. (Contributed by NM, 2-Aug-1994.)
 |-  ( A " B )  =  ran  ( A  |`  B )
 
Theoremxpeq1 4377 Equality theorem for cross product. (Contributed by NM, 4-Jul-1994.)
 |-  ( A  =  B  ->  ( A  X.  C )  =  ( B  X.  C ) )
 
Theoremxpeq2 4378 Equality theorem for cross product. (Contributed by NM, 5-Jul-1994.)
 |-  ( A  =  B  ->  ( C  X.  A )  =  ( C  X.  B ) )
 
Theoremelxpi 4379* Membership in a cross product. Uses fewer axioms than elxp 4380. (Contributed by NM, 4-Jul-1994.)
 |-  ( A  e.  ( B  X.  C )  ->  E. x E. y ( A  =  <. x ,  y >.  /\  ( x  e.  B  /\  y  e.  C ) ) )
 
Theoremelxp 4380* Membership in a cross product. (Contributed by NM, 4-Jul-1994.)
 |-  ( A  e.  ( B  X.  C )  <->  E. x E. y
 ( A  =  <. x ,  y >.  /\  ( x  e.  B  /\  y  e.  C )
 ) )
 
Theoremelxp2 4381* Membership in a cross product. (Contributed by NM, 23-Feb-2004.)
 |-  ( A  e.  ( B  X.  C )  <->  E. x  e.  B  E. y  e.  C  A  =  <. x ,  y >. )
 
Theoremxpeq12 4382 Equality theorem for cross product. (Contributed by FL, 31-Aug-2009.)
 |-  ( ( A  =  B  /\  C  =  D )  ->  ( A  X.  C )  =  ( B  X.  D ) )
 
Theoremxpeq1i 4383 Equality inference for cross product. (Contributed by NM, 21-Dec-2008.)
 |-  A  =  B   =>    |-  ( A  X.  C )  =  ( B  X.  C )
 
Theoremxpeq2i 4384 Equality inference for cross product. (Contributed by NM, 21-Dec-2008.)
 |-  A  =  B   =>    |-  ( C  X.  A )  =  ( C  X.  B )
 
Theoremxpeq12i 4385 Equality inference for cross product. (Contributed by FL, 31-Aug-2009.)
 |-  A  =  B   &    |-  C  =  D   =>    |-  ( A  X.  C )  =  ( B  X.  D )
 
Theoremxpeq1d 4386 Equality deduction for cross product. (Contributed by Jeff Madsen, 17-Jun-2010.)
 |-  ( ph  ->  A  =  B )   =>    |-  ( ph  ->  ( A  X.  C )  =  ( B  X.  C ) )
 
Theoremxpeq2d 4387 Equality deduction for cross product. (Contributed by Jeff Madsen, 17-Jun-2010.)
 |-  ( ph  ->  A  =  B )   =>    |-  ( ph  ->  ( C  X.  A )  =  ( C  X.  B ) )
 
Theoremxpeq12d 4388 Equality deduction for cross product. (Contributed by NM, 8-Dec-2013.)
 |-  ( ph  ->  A  =  B )   &    |-  ( ph  ->  C  =  D )   =>    |-  ( ph  ->  ( A  X.  C )  =  ( B  X.  D ) )
 
Theoremnfxp 4389 Bound-variable hypothesis builder for cross product. (Contributed by NM, 15-Sep-2003.) (Revised by Mario Carneiro, 15-Oct-2016.)
 |-  F/_ x A   &    |-  F/_ x B   =>    |-  F/_ x ( A  X.  B )
 
Theorem0nelxp 4390 The empty set is not a member of a cross product. (Contributed by NM, 2-May-1996.) (Revised by Mario Carneiro, 26-Apr-2015.)
 |- 
 -.  (/)  e.  ( A  X.  B )
 
Theorem0nelelxp 4391 A member of a cross product (ordered pair) doesn't contain the empty set. (Contributed by NM, 15-Dec-2008.)
 |-  ( C  e.  ( A  X.  B )  ->  -.  (/)  e.  C )
 
Theoremopelxp 4392 Ordered pair membership in a cross product. (Contributed by NM, 15-Nov-1994.) (Proof shortened by Andrew Salmon, 12-Aug-2011.) (Revised by Mario Carneiro, 26-Apr-2015.)
 |-  ( <. A ,  B >.  e.  ( C  X.  D )  <->  ( A  e.  C  /\  B  e.  D ) )
 
Theorembrxp 4393 Binary relation on a cross product. (Contributed by NM, 22-Apr-2004.)
 |-  ( A ( C  X.  D ) B  <-> 
 ( A  e.  C  /\  B  e.  D ) )
 
Theoremopelxpi 4394 Ordered pair membership in a cross product (implication). (Contributed by NM, 28-May-1995.)
 |-  ( ( A  e.  C  /\  B  e.  D )  ->  <. A ,  B >.  e.  ( C  X.  D ) )
 
Theoremopelxp1 4395 The first member of an ordered pair of classes in a cross product belongs to first cross product argument. (Contributed by NM, 28-May-2008.) (Revised by Mario Carneiro, 26-Apr-2015.)
 |-  ( <. A ,  B >.  e.  ( C  X.  D )  ->  A  e.  C )
 
Theoremopelxp2 4396 The second member of an ordered pair of classes in a cross product belongs to second cross product argument. (Contributed by Mario Carneiro, 26-Apr-2015.)
 |-  ( <. A ,  B >.  e.  ( C  X.  D )  ->  B  e.  D )
 
Theoremotelxp1 4397 The first member of an ordered triple of classes in a cross product belongs to first cross product argument. (Contributed by NM, 28-May-2008.)
 |-  ( <. <. A ,  B >. ,  C >.  e.  (
 ( R  X.  S )  X.  T )  ->  A  e.  R )
 
Theoremrabxp 4398* Membership in a class builder restricted to a cross product. (Contributed by NM, 20-Feb-2014.)
 |-  ( x  =  <. y ,  z >.  ->  ( ph 
 <->  ps ) )   =>    |-  { x  e.  ( A  X.  B )  |  ph }  =  { <. y ,  z >.  |  ( y  e.  A  /\  z  e.  B  /\  ps ) }
 
Theorembrrelex12 4399 A true binary relation on a relation implies the arguments are sets. (This is a property of our ordered pair definition.) (Contributed by Mario Carneiro, 26-Apr-2015.)
 |-  ( ( Rel  R  /\  A R B ) 
 ->  ( A  e.  _V  /\  B  e.  _V )
 )
 
Theorembrrelex 4400 A true binary relation on a relation implies the first argument is a set. (This is a property of our ordered pair definition.) (Contributed by NM, 18-May-2004.) (Revised by Mario Carneiro, 26-Apr-2015.)
 |-  ( ( Rel  R  /\  A R B ) 
 ->  A  e.  _V )
    < 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 >