Quot/QuotProd.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Tue, 26 Jan 2010 08:55:55 +0100
changeset 929 e812f58fd128
parent 928 44c92eaa4fad
child 931 0879d144aaa3
permissions -rw-r--r--
continued
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
922
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
     8
  "prod_rel R1 R2 = (\<lambda>(a,b) (c,d). R1 a c \<and> R2 b d)"
593
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     9
922
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    10
declare [[map * = (prod_fun, prod_rel)]]
593
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    11
922
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    12
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    13
lemma prod_equivp[quot_equiv]:
593
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    14
  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
    15
  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
    16
  shows "equivp (prod_rel R1 R2)"
928
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    17
  unfolding equivp_reflp_symp_transp reflp_def symp_def transp_def
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    18
  apply (auto simp add: equivp_reflp[OF a] equivp_reflp[OF b])
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    19
  apply (simp only: equivp_symp[OF a])
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    20
  apply (simp only: equivp_symp[OF b])
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    21
  using equivp_transp[OF a] apply blast
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    22
  using equivp_transp[OF b] apply blast
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    23
  done
593
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    24
922
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    25
lemma prod_quotient[quot_thm]:
593
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    26
  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
    27
  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
    28
  shows "Quotient (prod_rel R1 R2) (prod_fun Abs1 Abs2) (prod_fun Rep1 Rep2)"
928
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    29
  unfolding Quotient_def
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    30
  using q1 q2
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    31
  apply (simp add: Quotient_abs_rep Quotient_rel_rep)
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    32
  using Quotient_rel[OF q1] Quotient_rel[OF q2]
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    33
  by blast
593
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    34
922
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    35
lemma pair_rsp[quot_respect]:
593
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    36
  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
    37
  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
    38
  shows "(R1 ===> R2 ===> prod_rel R1 R2) Pair Pair"
928
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    39
  by simp
593
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    40
648
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    41
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
    42
  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
    43
  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
    44
  shows "(Rep1 ---> Rep2 ---> (prod_fun Abs1 Abs2)) Pair = Pair"
928
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    45
  apply (simp add: expand_fun_eq)
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    46
  apply (simp add: Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2])
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    47
  done
648
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    48
922
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    49
lemma fst_rsp[quot_respect]:
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    50
  assumes "Quotient R1 Abs1 Rep1"
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    51
  assumes "Quotient R2 Abs2 Rep2"
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    52
  shows "(prod_rel R1 R2 ===> R1) fst fst"
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    53
  by simp
648
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    54
922
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    55
lemma fst_prs[quot_preserve]:
593
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    56
  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
    57
  assumes q2: "Quotient R2 Abs2 Rep2"
922
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    58
  shows "(prod_fun Rep1 Rep2 ---> Abs1) fst = fst"
928
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    59
  apply (simp add: expand_fun_eq)
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    60
  apply (simp add: Quotient_abs_rep[OF q1])
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    61
  done
593
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    62
922
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    63
lemma snd_rsp[quot_respect]:
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    64
  assumes "Quotient R1 Abs1 Rep1"
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    65
  assumes "Quotient R2 Abs2 Rep2"
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    66
  shows "(prod_rel R1 R2 ===> R2) snd snd"
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    67
  by simp
928
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    68
922
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    69
lemma snd_prs[quot_preserve]:
593
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"
922
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    72
  shows "(prod_fun Rep1 Rep2 ---> Abs2) snd = snd"
928
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    73
  apply (simp add: expand_fun_eq)
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    74
  apply (simp add: Quotient_abs_rep[OF q2])
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    75
  done
593
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    76
928
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    77
lemma prod_fun_id[id_simps]:
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    78
  shows "prod_fun id id = id"
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    79
  by (simp add: prod_fun_def)
593
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    80
928
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    81
lemma prod_rel_eq[id_simps]:
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    82
  shows "(prod_rel (op =) (op =)) = (op =)"
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
    83
  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
    84
  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
    85
  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
    86
593
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    87
end