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

Theorem funprg 5268
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 1021 . . . 4 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → 𝐴𝑉)
2 simp2l 1023 . . . 4 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → 𝐶𝑋)
3 funsng 5264 . . . 4 ((𝐴𝑉𝐶𝑋) → Fun {⟨𝐴, 𝐶⟩})
41, 2, 3syl2anc 411 . . 3 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → Fun {⟨𝐴, 𝐶⟩})
5 simp1r 1022 . . . 4 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → 𝐵𝑊)
6 simp2r 1024 . . . 4 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → 𝐷𝑌)
7 funsng 5264 . . . 4 ((𝐵𝑊𝐷𝑌) → Fun {⟨𝐵, 𝐷⟩})
85, 6, 7syl2anc 411 . . 3 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → Fun {⟨𝐵, 𝐷⟩})
9 dmsnopg 5102 . . . . . 6 (𝐶𝑋 → dom {⟨𝐴, 𝐶⟩} = {𝐴})
102, 9syl 14 . . . . 5 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → dom {⟨𝐴, 𝐶⟩} = {𝐴})
11 dmsnopg 5102 . . . . . 6 (𝐷𝑌 → dom {⟨𝐵, 𝐷⟩} = {𝐵})
126, 11syl 14 . . . . 5 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → dom {⟨𝐵, 𝐷⟩} = {𝐵})
1310, 12ineq12d 3339 . . . 4 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → (dom {⟨𝐴, 𝐶⟩} ∩ dom {⟨𝐵, 𝐷⟩}) = ({𝐴} ∩ {𝐵}))
14 disjsn2 3657 . . . . 5 (𝐴𝐵 → ({𝐴} ∩ {𝐵}) = ∅)
15143ad2ant3 1020 . . . 4 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → ({𝐴} ∩ {𝐵}) = ∅)
1613, 15eqtrd 2210 . . 3 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → (dom {⟨𝐴, 𝐶⟩} ∩ dom {⟨𝐵, 𝐷⟩}) = ∅)
17 funun 5262 . . 3 (((Fun {⟨𝐴, 𝐶⟩} ∧ Fun {⟨𝐵, 𝐷⟩}) ∧ (dom {⟨𝐴, 𝐶⟩} ∩ dom {⟨𝐵, 𝐷⟩}) = ∅) → Fun ({⟨𝐴, 𝐶⟩} ∪ {⟨𝐵, 𝐷⟩}))
184, 8, 16, 17syl21anc 1237 . 2 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → Fun ({⟨𝐴, 𝐶⟩} ∪ {⟨𝐵, 𝐷⟩}))
19 df-pr 3601 . . 3 {⟨𝐴, 𝐶⟩, ⟨𝐵, 𝐷⟩} = ({⟨𝐴, 𝐶⟩} ∪ {⟨𝐵, 𝐷⟩})
2019funeqi 5239 . 2 (Fun {⟨𝐴, 𝐶⟩, ⟨𝐵, 𝐷⟩} ↔ Fun ({⟨𝐴, 𝐶⟩} ∪ {⟨𝐵, 𝐷⟩}))
2118, 20sylibr 134 1 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → Fun {⟨𝐴, 𝐶⟩, ⟨𝐵, 𝐷⟩})
Colors of variables: wff set class
Syntax hints:  wi 4  wa 104  w3a 978   = wceq 1353  wcel 2148  wne 2347  cun 3129  cin 3130  c0 3424  {csn 3594  {cpr 3595  cop 3597  dom cdm 4628  Fun wfun 5212
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 614  ax-in2 615  ax-io 709  ax-5 1447  ax-7 1448  ax-gen 1449  ax-ie1 1493  ax-ie2 1494  ax-8 1504  ax-10 1505  ax-11 1506  ax-i12 1507  ax-bndl 1509  ax-4 1510  ax-17 1526  ax-i9 1530  ax-ial 1534  ax-i5r 1535  ax-14 2151  ax-ext 2159  ax-sep 4123  ax-pow 4176  ax-pr 4211
This theorem depends on definitions:  df-bi 117  df-3an 980  df-tru 1356  df-fal 1359  df-nf 1461  df-sb 1763  df-eu 2029  df-mo 2030  df-clab 2164  df-cleq 2170  df-clel 2173  df-nfc 2308  df-ne 2348  df-ral 2460  df-rex 2461  df-v 2741  df-dif 3133  df-un 3135  df-in 3137  df-ss 3144  df-nul 3425  df-pw 3579  df-sn 3600  df-pr 3601  df-op 3603  df-br 4006  df-opab 4067  df-id 4295  df-xp 4634  df-rel 4635  df-cnv 4636  df-co 4637  df-dm 4638  df-fun 5220
This theorem is referenced by:  funtpg  5269  funpr  5270  fnprg  5273  2strbasg  12580  2stropg  12581
  Copyright terms: Public domain W3C validator