| Mathbox for Jim Kingdon | 
      
      
      < Previous  
      Next >
      
       Nearby theorems  | 
  ||
| Mirrors > Home > ILE Home > Th. List > Mathboxes > trirec0xor | GIF version | ||
| Description: Version of trirec0 15688 with exclusive-or.
 The definition of a discrete field is sometimes stated in terms of exclusive-or but as proved here, this is equivalent to inclusive-or because the two disjuncts cannot be simultaneously true. (Contributed by Jim Kingdon, 10-Jun-2024.)  | 
| Ref | Expression | 
|---|---|
| trirec0xor | ⊢ (∀𝑥 ∈ ℝ ∀𝑦 ∈ ℝ (𝑥 < 𝑦 ∨ 𝑥 = 𝑦 ∨ 𝑦 < 𝑥) ↔ ∀𝑥 ∈ ℝ (∃𝑧 ∈ ℝ (𝑥 · 𝑧) = 1 ⊻ 𝑥 = 0)) | 
| Step | Hyp | Ref | Expression | 
|---|---|---|---|
| 1 | trirec0 15688 | . 2 ⊢ (∀𝑥 ∈ ℝ ∀𝑦 ∈ ℝ (𝑥 < 𝑦 ∨ 𝑥 = 𝑦 ∨ 𝑦 < 𝑥) ↔ ∀𝑥 ∈ ℝ (∃𝑧 ∈ ℝ (𝑥 · 𝑧) = 1 ∨ 𝑥 = 0)) | |
| 2 | 1ne0 9058 | . . . . . . . 8 ⊢ 1 ≠ 0 | |
| 3 | 2 | nesymi 2413 | . . . . . . 7 ⊢ ¬ 0 = 1 | 
| 4 | simpr 110 | . . . . . . . . . . 11 ⊢ (((𝑥 · 𝑧) = 1 ∧ 𝑥 = 0) → 𝑥 = 0) | |
| 5 | 4 | oveq1d 5937 | . . . . . . . . . 10 ⊢ (((𝑥 · 𝑧) = 1 ∧ 𝑥 = 0) → (𝑥 · 𝑧) = (0 · 𝑧)) | 
| 6 | mul02lem2 8414 | . . . . . . . . . 10 ⊢ (𝑧 ∈ ℝ → (0 · 𝑧) = 0) | |
| 7 | 5, 6 | sylan9eqr 2251 | . . . . . . . . 9 ⊢ ((𝑧 ∈ ℝ ∧ ((𝑥 · 𝑧) = 1 ∧ 𝑥 = 0)) → (𝑥 · 𝑧) = 0) | 
| 8 | simprl 529 | . . . . . . . . 9 ⊢ ((𝑧 ∈ ℝ ∧ ((𝑥 · 𝑧) = 1 ∧ 𝑥 = 0)) → (𝑥 · 𝑧) = 1) | |
| 9 | 7, 8 | eqtr3d 2231 | . . . . . . . 8 ⊢ ((𝑧 ∈ ℝ ∧ ((𝑥 · 𝑧) = 1 ∧ 𝑥 = 0)) → 0 = 1) | 
| 10 | 9 | rexlimiva 2609 | . . . . . . 7 ⊢ (∃𝑧 ∈ ℝ ((𝑥 · 𝑧) = 1 ∧ 𝑥 = 0) → 0 = 1) | 
| 11 | 3, 10 | mto 663 | . . . . . 6 ⊢ ¬ ∃𝑧 ∈ ℝ ((𝑥 · 𝑧) = 1 ∧ 𝑥 = 0) | 
| 12 | r19.41v 2653 | . . . . . 6 ⊢ (∃𝑧 ∈ ℝ ((𝑥 · 𝑧) = 1 ∧ 𝑥 = 0) ↔ (∃𝑧 ∈ ℝ (𝑥 · 𝑧) = 1 ∧ 𝑥 = 0)) | |
| 13 | 11, 12 | mtbi 671 | . . . . 5 ⊢ ¬ (∃𝑧 ∈ ℝ (𝑥 · 𝑧) = 1 ∧ 𝑥 = 0) | 
| 14 | 13 | biantru 302 | . . . 4 ⊢ ((∃𝑧 ∈ ℝ (𝑥 · 𝑧) = 1 ∨ 𝑥 = 0) ↔ ((∃𝑧 ∈ ℝ (𝑥 · 𝑧) = 1 ∨ 𝑥 = 0) ∧ ¬ (∃𝑧 ∈ ℝ (𝑥 · 𝑧) = 1 ∧ 𝑥 = 0))) | 
| 15 | df-xor 1387 | . . . 4 ⊢ ((∃𝑧 ∈ ℝ (𝑥 · 𝑧) = 1 ⊻ 𝑥 = 0) ↔ ((∃𝑧 ∈ ℝ (𝑥 · 𝑧) = 1 ∨ 𝑥 = 0) ∧ ¬ (∃𝑧 ∈ ℝ (𝑥 · 𝑧) = 1 ∧ 𝑥 = 0))) | |
| 16 | 14, 15 | bitr4i 187 | . . 3 ⊢ ((∃𝑧 ∈ ℝ (𝑥 · 𝑧) = 1 ∨ 𝑥 = 0) ↔ (∃𝑧 ∈ ℝ (𝑥 · 𝑧) = 1 ⊻ 𝑥 = 0)) | 
| 17 | 16 | ralbii 2503 | . 2 ⊢ (∀𝑥 ∈ ℝ (∃𝑧 ∈ ℝ (𝑥 · 𝑧) = 1 ∨ 𝑥 = 0) ↔ ∀𝑥 ∈ ℝ (∃𝑧 ∈ ℝ (𝑥 · 𝑧) = 1 ⊻ 𝑥 = 0)) | 
| 18 | 1, 17 | bitri 184 | 1 ⊢ (∀𝑥 ∈ ℝ ∀𝑦 ∈ ℝ (𝑥 < 𝑦 ∨ 𝑥 = 𝑦 ∨ 𝑦 < 𝑥) ↔ ∀𝑥 ∈ ℝ (∃𝑧 ∈ ℝ (𝑥 · 𝑧) = 1 ⊻ 𝑥 = 0)) | 
| Colors of variables: wff set class | 
| Syntax hints: ¬ wn 3 ∧ wa 104 ↔ wb 105 ∨ wo 709 ∨ w3o 979 = wceq 1364 ⊻ wxo 1386 ∈ wcel 2167 ∀wral 2475 ∃wrex 2476 class class class wbr 4033 (class class class)co 5922 ℝcr 7878 0cc0 7879 1c1 7880 · cmul 7884 < clt 8061 | 
| This theorem was proved from axioms: ax-mp 5 ax-1 6 ax-2 7 ax-ia1 106 ax-ia2 107 ax-ia3 108 ax-in1 615 ax-in2 616 ax-io 710 ax-5 1461 ax-7 1462 ax-gen 1463 ax-ie1 1507 ax-ie2 1508 ax-8 1518 ax-10 1519 ax-11 1520 ax-i12 1521 ax-bndl 1523 ax-4 1524 ax-17 1540 ax-i9 1544 ax-ial 1548 ax-i5r 1549 ax-13 2169 ax-14 2170 ax-ext 2178 ax-sep 4151 ax-pow 4207 ax-pr 4242 ax-un 4468 ax-setind 4573 ax-cnex 7970 ax-resscn 7971 ax-1cn 7972 ax-1re 7973 ax-icn 7974 ax-addcl 7975 ax-addrcl 7976 ax-mulcl 7977 ax-mulrcl 7978 ax-addcom 7979 ax-mulcom 7980 ax-addass 7981 ax-mulass 7982 ax-distr 7983 ax-i2m1 7984 ax-0lt1 7985 ax-1rid 7986 ax-0id 7987 ax-rnegex 7988 ax-precex 7989 ax-cnre 7990 ax-pre-ltirr 7991 ax-pre-ltwlin 7992 ax-pre-lttrn 7993 ax-pre-apti 7994 ax-pre-ltadd 7995 ax-pre-mulgt0 7996 ax-pre-mulext 7997 | 
| This theorem depends on definitions: df-bi 117 df-3or 981 df-3an 982 df-tru 1367 df-fal 1370 df-xor 1387 df-nf 1475 df-sb 1777 df-eu 2048 df-mo 2049 df-clab 2183 df-cleq 2189 df-clel 2192 df-nfc 2328 df-ne 2368 df-nel 2463 df-ral 2480 df-rex 2481 df-reu 2482 df-rmo 2483 df-rab 2484 df-v 2765 df-sbc 2990 df-dif 3159 df-un 3161 df-in 3163 df-ss 3170 df-pw 3607 df-sn 3628 df-pr 3629 df-op 3631 df-uni 3840 df-br 4034 df-opab 4095 df-id 4328 df-po 4331 df-iso 4332 df-xp 4669 df-rel 4670 df-cnv 4671 df-co 4672 df-dm 4673 df-iota 5219 df-fun 5260 df-fv 5266 df-riota 5877 df-ov 5925 df-oprab 5926 df-mpo 5927 df-pnf 8063 df-mnf 8064 df-xr 8065 df-ltxr 8066 df-le 8067 df-sub 8199 df-neg 8200 df-reap 8602 df-ap 8609 df-div 8700 | 
| This theorem is referenced by: (None) | 
| Copyright terms: Public domain | W3C validator |