Attic/Quot/Quotient_Product.thy
author Christian Urban <urbanc@in.tum.de>
Tue, 28 Jun 2011 00:48:57 +0100
changeset 2913 bc86f5c3bc65
parent 1260 9df6144e281b
permissions -rw-r--r--
proved the fcb also for sets (no restriction yet)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 937
diff changeset
     1
(*  Title:      Quotient_Product.thy
935
c96e007b512f cleaning of QuotProd; a little cleaning of QuotList
Christian Urban <urbanc@in.tum.de>
parents: 925
diff changeset
     2
    Author:     Cezary Kaliszyk and Christian Urban
c96e007b512f cleaning of QuotProd; a little cleaning of QuotList
Christian Urban <urbanc@in.tum.de>
parents: 925
diff changeset
     3
*)
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 937
diff changeset
     4
theory Quotient_Product
1129
9a86f0ef6503 Notation available locally
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1128
diff changeset
     5
imports Quotient Quotient_Syntax
593
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     6
begin
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     7
937
Christian Urban <urbanc@in.tum.de>
parents: 936
diff changeset
     8
section {* Quotient infrastructure for the product type. *}
935
c96e007b512f cleaning of QuotProd; a little cleaning of QuotList
Christian Urban <urbanc@in.tum.de>
parents: 925
diff changeset
     9
593
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    10
fun
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    11
  prod_rel
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    12
where
935
c96e007b512f cleaning of QuotProd; a little cleaning of QuotList
Christian Urban <urbanc@in.tum.de>
parents: 925
diff changeset
    13
  "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
    14
922
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    15
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
    16
922
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    17
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    18
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
    19
  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
    20
  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
    21
  shows "equivp (prod_rel R1 R2)"
935
c96e007b512f cleaning of QuotProd; a little cleaning of QuotList
Christian Urban <urbanc@in.tum.de>
parents: 925
diff changeset
    22
  apply(rule equivpI)
c96e007b512f cleaning of QuotProd; a little cleaning of QuotList
Christian Urban <urbanc@in.tum.de>
parents: 925
diff changeset
    23
  unfolding reflp_def symp_def transp_def
c96e007b512f cleaning of QuotProd; a little cleaning of QuotList
Christian Urban <urbanc@in.tum.de>
parents: 925
diff changeset
    24
  apply(simp_all add: split_paired_all)
c96e007b512f cleaning of QuotProd; a little cleaning of QuotList
Christian Urban <urbanc@in.tum.de>
parents: 925
diff changeset
    25
  apply(blast intro: equivp_reflp[OF a] equivp_reflp[OF b])
c96e007b512f cleaning of QuotProd; a little cleaning of QuotList
Christian Urban <urbanc@in.tum.de>
parents: 925
diff changeset
    26
  apply(blast intro: equivp_symp[OF a] equivp_symp[OF b])
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 937
diff changeset
    27
  apply(blast intro: equivp_transp[OF a] equivp_transp[OF b])
928
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    28
  done
593
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    29
922
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    30
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
    31
  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
    32
  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
    33
  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
    34
  unfolding Quotient_def
935
c96e007b512f cleaning of QuotProd; a little cleaning of QuotList
Christian Urban <urbanc@in.tum.de>
parents: 925
diff changeset
    35
  apply(simp add: split_paired_all)
c96e007b512f cleaning of QuotProd; a little cleaning of QuotList
Christian Urban <urbanc@in.tum.de>
parents: 925
diff changeset
    36
  apply(simp add: Quotient_abs_rep[OF q1] Quotient_rel_rep[OF q1])
c96e007b512f cleaning of QuotProd; a little cleaning of QuotList
Christian Urban <urbanc@in.tum.de>
parents: 925
diff changeset
    37
  apply(simp add: Quotient_abs_rep[OF q2] Quotient_rel_rep[OF q2])
928
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    38
  using q1 q2
935
c96e007b512f cleaning of QuotProd; a little cleaning of QuotList
Christian Urban <urbanc@in.tum.de>
parents: 925
diff changeset
    39
  unfolding Quotient_def
c96e007b512f cleaning of QuotProd; a little cleaning of QuotList
Christian Urban <urbanc@in.tum.de>
parents: 925
diff changeset
    40
  apply(blast)
c96e007b512f cleaning of QuotProd; a little cleaning of QuotList
Christian Urban <urbanc@in.tum.de>
parents: 925
diff changeset
    41
  done
593
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    42
935
c96e007b512f cleaning of QuotProd; a little cleaning of QuotList
Christian Urban <urbanc@in.tum.de>
parents: 925
diff changeset
    43
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
    44
  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
    45
  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
    46
  shows "(R1 ===> R2 ===> prod_rel R1 R2) Pair Pair"
937
Christian Urban <urbanc@in.tum.de>
parents: 936
diff changeset
    47
  by simp
593
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    48
935
c96e007b512f cleaning of QuotProd; a little cleaning of QuotList
Christian Urban <urbanc@in.tum.de>
parents: 925
diff changeset
    49
lemma Pair_prs[quot_preserve]:
648
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    50
  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
    51
  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
    52
  shows "(Rep1 ---> Rep2 ---> (prod_fun Abs1 Abs2)) Pair = Pair"
935
c96e007b512f cleaning of QuotProd; a little cleaning of QuotList
Christian Urban <urbanc@in.tum.de>
parents: 925
diff changeset
    53
  apply(simp add: expand_fun_eq)
c96e007b512f cleaning of QuotProd; a little cleaning of QuotList
Christian Urban <urbanc@in.tum.de>
parents: 925
diff changeset
    54
  apply(simp add: Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2])
928
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    55
  done
648
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    56
922
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    57
lemma fst_rsp[quot_respect]:
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    58
  assumes "Quotient R1 Abs1 Rep1"
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    59
  assumes "Quotient R2 Abs2 Rep2"
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    60
  shows "(prod_rel R1 R2 ===> R1) fst fst"
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    61
  by simp
648
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    62
922
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    63
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
    64
  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
    65
  assumes q2: "Quotient R2 Abs2 Rep2"
922
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    66
  shows "(prod_fun Rep1 Rep2 ---> Abs1) fst = fst"
935
c96e007b512f cleaning of QuotProd; a little cleaning of QuotList
Christian Urban <urbanc@in.tum.de>
parents: 925
diff changeset
    67
  apply(simp add: expand_fun_eq)
c96e007b512f cleaning of QuotProd; a little cleaning of QuotList
Christian Urban <urbanc@in.tum.de>
parents: 925
diff changeset
    68
  apply(simp add: Quotient_abs_rep[OF q1])
928
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    69
  done
593
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    70
922
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    71
lemma snd_rsp[quot_respect]:
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    72
  assumes "Quotient R1 Abs1 Rep1"
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    73
  assumes "Quotient R2 Abs2 Rep2"
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    74
  shows "(prod_rel R1 R2 ===> R2) snd snd"
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    75
  by simp
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 937
diff changeset
    76
922
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    77
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
    78
  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
    79
  assumes q2: "Quotient R2 Abs2 Rep2"
922
a2589b4bc442 tuned proofs (mainly in QuotProd)
Christian Urban <urbanc@in.tum.de>
parents: 913
diff changeset
    80
  shows "(prod_fun Rep1 Rep2 ---> Abs2) snd = snd"
935
c96e007b512f cleaning of QuotProd; a little cleaning of QuotList
Christian Urban <urbanc@in.tum.de>
parents: 925
diff changeset
    81
  apply(simp add: expand_fun_eq)
c96e007b512f cleaning of QuotProd; a little cleaning of QuotList
Christian Urban <urbanc@in.tum.de>
parents: 925
diff changeset
    82
  apply(simp add: Quotient_abs_rep[OF q2])
928
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    83
  done
593
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    84
931
0879d144aaa3 Generalized split_prs and split_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 928
diff changeset
    85
lemma split_rsp[quot_respect]:
936
Christian Urban <urbanc@in.tum.de>
parents: 935 931
diff changeset
    86
  shows "((R1 ===> R2 ===> (op =)) ===> (prod_rel R1 R2) ===> (op =)) split split"
931
0879d144aaa3 Generalized split_prs and split_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 928
diff changeset
    87
  by auto
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 937
diff changeset
    88
931
0879d144aaa3 Generalized split_prs and split_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 928
diff changeset
    89
lemma split_prs[quot_preserve]:
0879d144aaa3 Generalized split_prs and split_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 928
diff changeset
    90
  assumes q1: "Quotient R1 Abs1 Rep1"
0879d144aaa3 Generalized split_prs and split_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 928
diff changeset
    91
  and     q2: "Quotient R2 Abs2 Rep2"
0879d144aaa3 Generalized split_prs and split_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 928
diff changeset
    92
  shows "(((Abs1 ---> Abs2 ---> id) ---> prod_fun Rep1 Rep2 ---> id) split) = split"
0879d144aaa3 Generalized split_prs and split_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 928
diff changeset
    93
  by (simp add: expand_fun_eq Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2])
0879d144aaa3 Generalized split_prs and split_rsp
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 928
diff changeset
    94
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 937
diff changeset
    95
lemma prod_fun_id[id_simps]:
928
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    96
  shows "prod_fun id id = id"
44c92eaa4fad More eqreflection/equiv cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 925
diff changeset
    97
  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
    98
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 937
diff changeset
    99
lemma prod_rel_eq[id_simps]:
936
Christian Urban <urbanc@in.tum.de>
parents: 935 931
diff changeset
   100
  shows "prod_rel (op =) (op =) = (op =)"
Christian Urban <urbanc@in.tum.de>
parents: 935 931
diff changeset
   101
  by (simp add: expand_fun_eq)
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 937
diff changeset
   102
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
   103
593
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   104
end