Quot/QuotProd.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Fri, 08 Jan 2010 15:02:12 +0100
changeset 829 42b90994ac77
parent 699 aa157e957655
child 867 9e247b9505f0
permissions -rw-r--r--
map and rel simps for all quotients; needed when changing the relations to aggregate ones.
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
648
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
     2
imports QuotMain
593
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
699
aa157e957655 added maps-printout and tuned some comments
Christian Urban <urbanc@in.tum.de>
parents: 668
diff changeset
    29
using q1 q2
aa157e957655 added maps-printout and tuned some comments
Christian Urban <urbanc@in.tum.de>
parents: 668
diff changeset
    30
apply (simp add: Quotient_abs_rep Quotient_abs_rep Quotient_rel_rep Quotient_rel_rep)
593
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    31
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
    32
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    33
lemma pair_rsp:
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    34
  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
    35
  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
    36
  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
    37
by auto
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    38
648
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    39
lemma pair_prs_aux:
593
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    40
  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
    41
  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
    42
  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
    43
  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
    44
648
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    45
lemma pair_prs[quot_preserve]:
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    46
  assumes q1: "Quotient R1 Abs1 Rep1"
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    47
  assumes q2: "Quotient R2 Abs2 Rep2"
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    48
  shows "(Rep1 ---> Rep2 ---> (prod_fun Abs1 Abs2)) Pair = Pair"
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    49
apply(simp only: expand_fun_eq fun_map.simps pair_prs_aux[OF q1 q2])
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    50
apply(simp)
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    51
done
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    52
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    53
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    54
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    55
593
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    56
(* 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
    57
(* 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
    58
lemma fst_rsp:
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    59
  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
    60
  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
    61
  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
    62
  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
    63
  using a
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 p1)
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    65
  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
    66
  apply(auto)
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    67
  done
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    68
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    69
lemma snd_rsp:
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    70
  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
    71
  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
    72
  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
    73
  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
    74
  using a
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    75
  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
    76
  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
    77
  apply(auto)
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    78
  done
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
lemma fst_prs:
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    81
  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
    82
  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
    83
  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
    84
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
    85
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    86
lemma snd_prs:
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    87
  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
    88
  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
    89
  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
    90
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
    91
668
ef5b941f00e2 Code cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 648
diff changeset
    92
lemma prod_fun_id[id_simps]: "prod_fun id id \<equiv> id"
ef5b941f00e2 Code cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 648
diff changeset
    93
  by (rule eq_reflection) (simp add: prod_fun_def)
648
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    94
829
42b90994ac77 map and rel simps for all quotients; needed when changing the relations to aggregate ones.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 699
diff changeset
    95
lemma prod_rel_eq[id_simps]: "prod_rel op = op = \<equiv> op ="
42b90994ac77 map and rel simps for all quotients; needed when changing the relations to aggregate ones.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 699
diff changeset
    96
  apply (rule eq_reflection)
42b90994ac77 map and rel simps for all quotients; needed when changing the relations to aggregate ones.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 699
diff changeset
    97
  apply (rule ext)+
42b90994ac77 map and rel simps for all quotients; needed when changing the relations to aggregate ones.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 699
diff changeset
    98
  apply auto
42b90994ac77 map and rel simps for all quotients; needed when changing the relations to aggregate ones.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 699
diff changeset
    99
  done
42b90994ac77 map and rel simps for all quotients; needed when changing the relations to aggregate ones.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 699
diff changeset
   100
648
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
   101
section {* general setup for products *}
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
   102
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
   103
declare [[map * = (prod_fun, prod_rel)]]
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
   104
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
   105
lemmas [quot_thm] = prod_quotient
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
   106
lemmas [quot_respect] = pair_rsp
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
   107
lemmas [quot_equiv] = prod_equivp
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
   108
593
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   109
end