ILE Home Intuitionistic Logic Explorer < Previous   Next >
Nearby theorems
Mirrors  >  Home  >  ILE Home  >  Th. List  >  funprg GIF version

Theorem funprg 5248
Description: A set of two pairs is a function if their first members are different. (Contributed by FL, 26-Jun-2011.)
Assertion
Ref Expression
funprg (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → Fun {⟨𝐴, 𝐶⟩, ⟨𝐵, 𝐷⟩})

Proof of Theorem funprg
StepHypRef Expression
1 simp1l 1016 . . . 4 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → 𝐴𝑉)
2 simp2l 1018 . . . 4 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → 𝐶𝑋)
3 funsng 5244 . . . 4 ((𝐴𝑉𝐶𝑋) → Fun {⟨𝐴, 𝐶⟩})
41, 2, 3syl2anc 409 . . 3 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → Fun {⟨𝐴, 𝐶⟩})
5 simp1r 1017 . . . 4 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → 𝐵𝑊)
6 simp2r 1019 . . . 4 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → 𝐷𝑌)
7 funsng 5244 . . . 4 ((𝐵𝑊𝐷𝑌) → Fun {⟨𝐵, 𝐷⟩})
85, 6, 7syl2anc 409 . . 3 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → Fun {⟨𝐵, 𝐷⟩})
9 dmsnopg 5082 . . . . . 6 (𝐶𝑋 → dom {⟨𝐴, 𝐶⟩} = {𝐴})
102, 9syl 14 . . . . 5 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → dom {⟨𝐴, 𝐶⟩} = {𝐴})
11 dmsnopg 5082 . . . . . 6 (𝐷𝑌 → dom {⟨𝐵, 𝐷⟩} = {𝐵})
126, 11syl 14 . . . . 5 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → dom {⟨𝐵, 𝐷⟩} = {𝐵})
1310, 12ineq12d 3329 . . . 4 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → (dom {⟨𝐴, 𝐶⟩} ∩ dom {⟨𝐵, 𝐷⟩}) = ({𝐴} ∩ {𝐵}))
14 disjsn2 3646 . . . . 5 (𝐴𝐵 → ({𝐴} ∩ {𝐵}) = ∅)
15143ad2ant3 1015 . . . 4 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → ({𝐴} ∩ {𝐵}) = ∅)
1613, 15eqtrd 2203 . . 3 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → (dom {⟨𝐴, 𝐶⟩} ∩ dom {⟨𝐵, 𝐷⟩}) = ∅)
17 funun 5242 . . 3 (((Fun {⟨𝐴, 𝐶⟩} ∧ Fun {⟨𝐵, 𝐷⟩}) ∧ (dom {⟨𝐴, 𝐶⟩} ∩ dom {⟨𝐵, 𝐷⟩}) = ∅) → Fun ({⟨𝐴, 𝐶⟩} ∪ {⟨𝐵, 𝐷⟩}))
184, 8, 16, 17syl21anc 1232 . 2 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → Fun ({⟨𝐴, 𝐶⟩} ∪ {⟨𝐵, 𝐷⟩}))
19 df-pr 3590 . . 3 {⟨𝐴, 𝐶⟩, ⟨𝐵, 𝐷⟩} = ({⟨𝐴, 𝐶⟩} ∪ {⟨𝐵, 𝐷⟩})
2019funeqi 5219 . 2 (Fun {⟨𝐴, 𝐶⟩, ⟨𝐵, 𝐷⟩} ↔ Fun ({⟨𝐴, 𝐶⟩} ∪ {⟨𝐵, 𝐷⟩}))
2118, 20sylibr 133 1 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → Fun {⟨𝐴, 𝐶⟩, ⟨𝐵, 𝐷⟩})
Colors of variables: wff set class
Syntax hints:  wi 4  wa 103  w3a 973   = wceq 1348  wcel 2141  wne 2340  cun 3119  cin 3120  c0 3414  {csn 3583  {cpr 3584  cop 3586  dom cdm 4611  Fun wfun 5192
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-ia1 105  ax-ia2 106  ax-ia3 107  ax-in1 609  ax-in2 610  ax-io 704  ax-5 1440  ax-7 1441  ax-gen 1442  ax-ie1 1486  ax-ie2 1487  ax-8 1497  ax-10 1498  ax-11 1499  ax-i12 1500  ax-bndl 1502  ax-4 1503  ax-17 1519  ax-i9 1523  ax-ial 1527  ax-i5r 1528  ax-14 2144  ax-ext 2152  ax-sep 4107  ax-pow 4160  ax-pr 4194
This theorem depends on definitions:  df-bi 116  df-3an 975  df-tru 1351  df-fal 1354  df-nf 1454  df-sb 1756  df-eu 2022  df-mo 2023  df-clab 2157  df-cleq 2163  df-clel 2166  df-nfc 2301  df-ne 2341  df-ral 2453  df-rex 2454  df-v 2732  df-dif 3123  df-un 3125  df-in 3127  df-ss 3134  df-nul 3415  df-pw 3568  df-sn 3589  df-pr 3590  df-op 3592  df-br 3990  df-opab 4051  df-id 4278  df-xp 4617  df-rel 4618  df-cnv 4619  df-co 4620  df-dm 4621  df-fun 5200
This theorem is referenced by:  funtpg  5249  funpr  5250  fnprg  5253  2strbasg  12519  2stropg  12520
  Copyright terms: Public domain W3C validator