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

Theorem funprg 5304
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 1023 . . . 4 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → 𝐴𝑉)
2 simp2l 1025 . . . 4 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → 𝐶𝑋)
3 funsng 5300 . . . 4 ((𝐴𝑉𝐶𝑋) → Fun {⟨𝐴, 𝐶⟩})
41, 2, 3syl2anc 411 . . 3 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → Fun {⟨𝐴, 𝐶⟩})
5 simp1r 1024 . . . 4 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → 𝐵𝑊)
6 simp2r 1026 . . . 4 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → 𝐷𝑌)
7 funsng 5300 . . . 4 ((𝐵𝑊𝐷𝑌) → Fun {⟨𝐵, 𝐷⟩})
85, 6, 7syl2anc 411 . . 3 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → Fun {⟨𝐵, 𝐷⟩})
9 dmsnopg 5137 . . . . . 6 (𝐶𝑋 → dom {⟨𝐴, 𝐶⟩} = {𝐴})
102, 9syl 14 . . . . 5 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → dom {⟨𝐴, 𝐶⟩} = {𝐴})
11 dmsnopg 5137 . . . . . 6 (𝐷𝑌 → dom {⟨𝐵, 𝐷⟩} = {𝐵})
126, 11syl 14 . . . . 5 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → dom {⟨𝐵, 𝐷⟩} = {𝐵})
1310, 12ineq12d 3361 . . . 4 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → (dom {⟨𝐴, 𝐶⟩} ∩ dom {⟨𝐵, 𝐷⟩}) = ({𝐴} ∩ {𝐵}))
14 disjsn2 3681 . . . . 5 (𝐴𝐵 → ({𝐴} ∩ {𝐵}) = ∅)
15143ad2ant3 1022 . . . 4 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → ({𝐴} ∩ {𝐵}) = ∅)
1613, 15eqtrd 2226 . . 3 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → (dom {⟨𝐴, 𝐶⟩} ∩ dom {⟨𝐵, 𝐷⟩}) = ∅)
17 funun 5298 . . 3 (((Fun {⟨𝐴, 𝐶⟩} ∧ Fun {⟨𝐵, 𝐷⟩}) ∧ (dom {⟨𝐴, 𝐶⟩} ∩ dom {⟨𝐵, 𝐷⟩}) = ∅) → Fun ({⟨𝐴, 𝐶⟩} ∪ {⟨𝐵, 𝐷⟩}))
184, 8, 16, 17syl21anc 1248 . 2 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → Fun ({⟨𝐴, 𝐶⟩} ∪ {⟨𝐵, 𝐷⟩}))
19 df-pr 3625 . . 3 {⟨𝐴, 𝐶⟩, ⟨𝐵, 𝐷⟩} = ({⟨𝐴, 𝐶⟩} ∪ {⟨𝐵, 𝐷⟩})
2019funeqi 5275 . 2 (Fun {⟨𝐴, 𝐶⟩, ⟨𝐵, 𝐷⟩} ↔ Fun ({⟨𝐴, 𝐶⟩} ∪ {⟨𝐵, 𝐷⟩}))
2118, 20sylibr 134 1 (((𝐴𝑉𝐵𝑊) ∧ (𝐶𝑋𝐷𝑌) ∧ 𝐴𝐵) → Fun {⟨𝐴, 𝐶⟩, ⟨𝐵, 𝐷⟩})
Colors of variables: wff set class
Syntax hints:  wi 4  wa 104  w3a 980   = wceq 1364  wcel 2164  wne 2364  cun 3151  cin 3152  c0 3446  {csn 3618  {cpr 3619  cop 3621  dom cdm 4659  Fun wfun 5248
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 1458  ax-7 1459  ax-gen 1460  ax-ie1 1504  ax-ie2 1505  ax-8 1515  ax-10 1516  ax-11 1517  ax-i12 1518  ax-bndl 1520  ax-4 1521  ax-17 1537  ax-i9 1541  ax-ial 1545  ax-i5r 1546  ax-14 2167  ax-ext 2175  ax-sep 4147  ax-pow 4203  ax-pr 4238
This theorem depends on definitions:  df-bi 117  df-3an 982  df-tru 1367  df-fal 1370  df-nf 1472  df-sb 1774  df-eu 2045  df-mo 2046  df-clab 2180  df-cleq 2186  df-clel 2189  df-nfc 2325  df-ne 2365  df-ral 2477  df-rex 2478  df-v 2762  df-dif 3155  df-un 3157  df-in 3159  df-ss 3166  df-nul 3447  df-pw 3603  df-sn 3624  df-pr 3625  df-op 3627  df-br 4030  df-opab 4091  df-id 4324  df-xp 4665  df-rel 4666  df-cnv 4667  df-co 4668  df-dm 4669  df-fun 5256
This theorem is referenced by:  funtpg  5305  funpr  5306  fnprg  5309  2strbasg  12737  2stropg  12738
  Copyright terms: Public domain W3C validator