MPE Home Metamath Proof Explorer < Previous   Next >
Nearby theorems
Mirrors  >  Home  >  MPE Home  >  Th. List  >  df-if Structured version   Visualization version   GIF version

Definition df-if 4037
Description: Define the conditional operator. Read if(𝜑, 𝐴, 𝐵) as "if 𝜑 then 𝐴 else 𝐵." See iftrue 4042 and iffalse 4045 for its values. In mathematical literature, this operator is rarely defined formally but is implicit in informal definitions such as "let f(x)=0 if x=0 and 1/x otherwise." (In older versions of this database, this operator was denoted "ded" and called the "deduction class.")

An important use for us is in conjunction with the weak deduction theorem, which converts a hypothesis into an antecedent. In that role, 𝐴 is a class variable in the hypothesis and 𝐵 is a class (usually a constant) that makes the hypothesis true when it is substituted for 𝐴. See dedth 4089 for the main part of the weak deduction theorem, elimhyp 4096 to eliminate a hypothesis, and keephyp 4102 to keep a hypothesis. See the Deduction Theorem link on the Metamath Proof Explorer Home Page for a description of the weak deduction theorem. (Contributed by NM, 15-May-1999.)

Assertion
Ref Expression
df-if if(𝜑, 𝐴, 𝐵) = {𝑥 ∣ ((𝑥𝐴𝜑) ∨ (𝑥𝐵 ∧ ¬ 𝜑))}
Distinct variable groups:   𝜑,𝑥   𝑥,𝐴   𝑥,𝐵

Detailed syntax breakdown of Definition df-if
StepHypRef Expression
1 wph . . 3 wff 𝜑
2 cA . . 3 class 𝐴
3 cB . . 3 class 𝐵
41, 2, 3cif 4036 . 2 class if(𝜑, 𝐴, 𝐵)
5 vx . . . . . . 7 setvar 𝑥
65cv 1474 . . . . . 6 class 𝑥
76, 2wcel 1977 . . . . 5 wff 𝑥𝐴
87, 1wa 383 . . . 4 wff (𝑥𝐴𝜑)
96, 3wcel 1977 . . . . 5 wff 𝑥𝐵
101wn 3 . . . . 5 wff ¬ 𝜑
119, 10wa 383 . . . 4 wff (𝑥𝐵 ∧ ¬ 𝜑)
128, 11wo 382 . . 3 wff ((𝑥𝐴𝜑) ∨ (𝑥𝐵 ∧ ¬ 𝜑))
1312, 5cab 2596 . 2 class {𝑥 ∣ ((𝑥𝐴𝜑) ∨ (𝑥𝐵 ∧ ¬ 𝜑))}
144, 13wceq 1475 1 wff if(𝜑, 𝐴, 𝐵) = {𝑥 ∣ ((𝑥𝐴𝜑) ∨ (𝑥𝐵 ∧ ¬ 𝜑))}
Colors of variables: wff setvar class
This definition is referenced by:  dfif2  4038  dfif6  4039  iffalse  4045  rabsnifsb  4201  bj-dfifc2  31734
  Copyright terms: Public domain W3C validator