Quot/QuotProd.thy
author Christian Urban <urbanc@in.tum.de>
Mon, 07 Dec 2009 14:09:50 +0100
changeset 597 8a1c8dc72b5c
parent 593 QuotProd.thy@18eac4596ef1
child 648 830b58c2fa94
permissions -rw-r--r--
directory re-arrangement
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
593
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     1
theory QuotProd
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     2
imports QuotScript
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     3
begin
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     4
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     5
fun
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     6
  prod_rel
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     7
where
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     8
  "prod_rel r1 r2 = (\<lambda>(a,b) (c,d). r1 a c \<and> r2 b d)"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     9
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    10
(* prod_fun is a good mapping function *)
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    11
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    12
lemma prod_equivp:
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    13
  assumes a: "equivp R1"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    14
  assumes b: "equivp R2"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    15
  shows "equivp (prod_rel R1 R2)"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    16
unfolding equivp_reflp_symp_transp reflp_def symp_def transp_def
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    17
apply(auto simp add: equivp_reflp[OF a] equivp_reflp[OF b])
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    18
apply(simp only: equivp_symp[OF a])
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    19
apply(simp only: equivp_symp[OF b])
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    20
using equivp_transp[OF a] apply blast
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    21
using equivp_transp[OF b] apply blast
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    22
done
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    23
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    24
lemma prod_quotient:
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    25
  assumes q1: "Quotient R1 Abs1 Rep1"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    26
  assumes q2: "Quotient R2 Abs2 Rep2"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    27
  shows "Quotient (prod_rel R1 R2) (prod_fun Abs1 Abs2) (prod_fun Rep1 Rep2)"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    28
unfolding Quotient_def
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    29
apply (simp add: Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2] Quotient_rel_rep[OF q1] Quotient_rel_rep[OF q2])
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    30
using Quotient_rel[OF q1] Quotient_rel[OF q2] by blast
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    31
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    32
lemma pair_rsp:
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    33
  assumes q1: "Quotient R1 Abs1 Rep1"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    34
  assumes q2: "Quotient R2 Abs2 Rep2"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    35
  shows "(R1 ===> R2 ===> prod_rel R1 R2) Pair Pair"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    36
by auto
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    37
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    38
lemma pair_prs:
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    39
  assumes q1: "Quotient R1 Abs1 Rep1"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    40
  assumes q2: "Quotient R2 Abs2 Rep2"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    41
  shows "(prod_fun Abs1 Abs2) (Rep1 l, Rep2 r) \<equiv> (l, r)"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    42
  by (simp add: Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2])
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    43
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    44
(* TODO: Is the quotient assumption q1 necessary? *)
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    45
(* TODO: Aren't there hard to use later? *)
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    46
lemma fst_rsp:
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    47
  assumes q1: "Quotient R1 Abs1 Rep1"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    48
  assumes q2: "Quotient R2 Abs2 Rep2"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    49
  assumes a: "(prod_rel R1 R2) p1 p2"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    50
  shows "R1 (fst p1) (fst p2)"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    51
  using a
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    52
  apply(case_tac p1)
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    53
  apply(case_tac p2)
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    54
  apply(auto)
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    55
  done
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    56
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    57
lemma snd_rsp:
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    58
  assumes q1: "Quotient R1 Abs1 Rep1"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    59
  assumes q2: "Quotient R2 Abs2 Rep2"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    60
  assumes a: "(prod_rel R1 R2) p1 p2"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    61
  shows "R2 (snd p1) (snd p2)"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    62
  using a
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    63
  apply(case_tac p1)
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    64
  apply(case_tac p2)
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    65
  apply(auto)
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    66
  done
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    67
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    68
lemma fst_prs:
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    69
  assumes q1: "Quotient R1 Abs1 Rep1"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    70
  assumes q2: "Quotient R2 Abs2 Rep2"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    71
  shows "Abs1 (fst ((prod_fun Rep1 Rep2) p)) = fst p"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    72
by (case_tac p) (auto simp add: Quotient_abs_rep[OF q1])
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    73
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    74
lemma snd_prs:
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    75
  assumes q1: "Quotient R1 Abs1 Rep1"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    76
  assumes q2: "Quotient R2 Abs2 Rep2"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    77
  shows "Abs2 (snd ((prod_fun Rep1 Rep2) p)) = snd p"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    78
by (case_tac p) (auto simp add: Quotient_abs_rep[OF q2])
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    79
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    80
end