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

Theorem funprg 5238
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 1011 . . . 4 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → 𝐴𝑉)
2 simp2l 1013 . . . 4 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → 𝐶𝑋)
3 funsng 5234 . . . 4 ((𝐴𝑉𝐶𝑋) → Fun {⟨𝐴, 𝐶⟩})
41, 2, 3syl2anc 409 . . 3 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → Fun {⟨𝐴, 𝐶⟩})
5 simp1r 1012 . . . 4 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → 𝐵𝑊)
6 simp2r 1014 . . . 4 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → 𝐷𝑌)
7 funsng 5234 . . . 4 ((𝐵𝑊𝐷𝑌) → Fun {⟨𝐵, 𝐷⟩})
85, 6, 7syl2anc 409 . . 3 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → Fun {⟨𝐵, 𝐷⟩})
9 dmsnopg 5075 . . . . . 6 (𝐶𝑋 → dom {⟨𝐴, 𝐶⟩} = {𝐴})
102, 9syl 14 . . . . 5 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → dom {⟨𝐴, 𝐶⟩} = {𝐴})
11 dmsnopg 5075 . . . . . 6 (𝐷𝑌 → dom {⟨𝐵, 𝐷⟩} = {𝐵})
126, 11syl 14 . . . . 5 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → dom {⟨𝐵, 𝐷⟩} = {𝐵})
1310, 12ineq12d 3324 . . . 4 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → (dom {⟨𝐴, 𝐶⟩} ∩ dom {⟨𝐵, 𝐷⟩}) = ({𝐴} ∩ {𝐵}))
14 disjsn2 3639 . . . . 5 (𝐴𝐵 → ({𝐴} ∩ {𝐵}) = ∅)
15143ad2ant3 1010 . . . 4 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → ({𝐴} ∩ {𝐵}) = ∅)
1613, 15eqtrd 2198 . . 3 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → (dom {⟨𝐴, 𝐶⟩} ∩ dom {⟨𝐵, 𝐷⟩}) = ∅)
17 funun 5232 . . 3 (((Fun {⟨𝐴, 𝐶⟩} ∧ Fun {⟨𝐵, 𝐷⟩}) ∧ (dom {⟨𝐴, 𝐶⟩} ∩ dom {⟨𝐵, 𝐷⟩}) = ∅) → Fun ({⟨𝐴, 𝐶⟩} ∪ {⟨𝐵, 𝐷⟩}))
184, 8, 16, 17syl21anc 1227 . 2 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → Fun ({⟨𝐴, 𝐶⟩} ∪ {⟨𝐵, 𝐷⟩}))
19 df-pr 3583 . . 3 {⟨𝐴, 𝐶⟩, ⟨𝐵, 𝐷⟩} = ({⟨𝐴, 𝐶⟩} ∪ {⟨𝐵, 𝐷⟩})
2019funeqi 5209 . 2 (Fun {⟨𝐴, 𝐶⟩, ⟨𝐵, 𝐷⟩} ↔ Fun ({⟨𝐴, 𝐶⟩} ∪ {⟨𝐵, 𝐷⟩}))
2118, 20sylibr 133 1 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → Fun {⟨𝐴, 𝐶⟩, ⟨𝐵, 𝐷⟩})
Colors of variables: wff set class
Syntax hints:  wi 4  wa 103  w3a 968   = wceq 1343  wcel 2136  wne 2336  cun 3114  cin 3115  c0 3409  {csn 3576  {cpr 3577  cop 3579  dom cdm 4604  Fun wfun 5182
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 604  ax-in2 605  ax-io 699  ax-5 1435  ax-7 1436  ax-gen 1437  ax-ie1 1481  ax-ie2 1482  ax-8 1492  ax-10 1493  ax-11 1494  ax-i12 1495  ax-bndl 1497  ax-4 1498  ax-17 1514  ax-i9 1518  ax-ial 1522  ax-i5r 1523  ax-14 2139  ax-ext 2147  ax-sep 4100  ax-pow 4153  ax-pr 4187
This theorem depends on definitions:  df-bi 116  df-3an 970  df-tru 1346  df-fal 1349  df-nf 1449  df-sb 1751  df-eu 2017  df-mo 2018  df-clab 2152  df-cleq 2158  df-clel 2161  df-nfc 2297  df-ne 2337  df-ral 2449  df-rex 2450  df-v 2728  df-dif 3118  df-un 3120  df-in 3122  df-ss 3129  df-nul 3410  df-pw 3561  df-sn 3582  df-pr 3583  df-op 3585  df-br 3983  df-opab 4044  df-id 4271  df-xp 4610  df-rel 4611  df-cnv 4612  df-co 4613  df-dm 4614  df-fun 5190
This theorem is referenced by:  funtpg  5239  funpr  5240  fnprg  5243  2strbasg  12496  2stropg  12497
  Copyright terms: Public domain W3C validator