Attic/Quot/Quotient_Sum.thy
author Christian Urban <urbanc@in.tum.de>
Fri, 14 May 2010 17:40:43 +0100
changeset 2136 2fc55508a6d0
parent 1260 9df6144e281b
permissions -rw-r--r--
polished example
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: 1122
diff changeset
     1
(*  Title:      Quotient_Sum.thy
934
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
     2
    Author:     Cezary Kaliszyk and Christian Urban
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
     3
*)
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1122
diff changeset
     4
theory Quotient_Sum
1129
9a86f0ef6503 Notation available locally
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1128
diff changeset
     5
imports Quotient Quotient_Syntax
0
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
begin
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
937
Christian Urban <urbanc@in.tum.de>
parents: 936
diff changeset
     8
section {* Quotient infrastructure for the sum type. *}
Christian Urban <urbanc@in.tum.de>
parents: 936
diff changeset
     9
0
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
fun
545
95371a8b17e1 Cleaning the Quotients file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 538
diff changeset
    11
  sum_rel
0
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
where
545
95371a8b17e1 Cleaning the Quotients file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 538
diff changeset
    13
  "sum_rel R1 R2 (Inl a1) (Inl b1) = R1 a1 b1"
95371a8b17e1 Cleaning the Quotients file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 538
diff changeset
    14
| "sum_rel R1 R2 (Inl a1) (Inr b2) = False"
95371a8b17e1 Cleaning the Quotients file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 538
diff changeset
    15
| "sum_rel R1 R2 (Inr a2) (Inl b1) = False"
95371a8b17e1 Cleaning the Quotients file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 538
diff changeset
    16
| "sum_rel R1 R2 (Inr a2) (Inr b2) = R2 a2 b2"
0
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
fun
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
  sum_map
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
where
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
  "sum_map f1 f2 (Inl a) = Inl (f1 a)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
| "sum_map f1 f2 (Inr a) = Inr (f2 a)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 698
diff changeset
    24
declare [[map "+" = (sum_map, sum_rel)]]
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 698
diff changeset
    25
545
95371a8b17e1 Cleaning the Quotients file
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 538
diff changeset
    26
934
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    27
text {* should probably be in Sum_Type.thy *}
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1122
diff changeset
    28
lemma split_sum_all:
934
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    29
  shows "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x. P (Inl x)) \<and> (\<forall>x. P (Inr x))"
937
Christian Urban <urbanc@in.tum.de>
parents: 936
diff changeset
    30
  apply(auto)
Christian Urban <urbanc@in.tum.de>
parents: 936
diff changeset
    31
  apply(case_tac x)
Christian Urban <urbanc@in.tum.de>
parents: 936
diff changeset
    32
  apply(simp_all)
Christian Urban <urbanc@in.tum.de>
parents: 936
diff changeset
    33
  done
934
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    34
698
ed44eaaea63e Option and Sum quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 613
diff changeset
    35
lemma sum_equivp[quot_equiv]:
ed44eaaea63e Option and Sum quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 613
diff changeset
    36
  assumes a: "equivp R1"
ed44eaaea63e Option and Sum quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 613
diff changeset
    37
  assumes b: "equivp R2"
ed44eaaea63e Option and Sum quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 613
diff changeset
    38
  shows "equivp (sum_rel R1 R2)"
924
5455b19ef138 re-inserted lemma in QuotList
Christian Urban <urbanc@in.tum.de>
parents: 829
diff changeset
    39
  apply(rule equivpI)
5455b19ef138 re-inserted lemma in QuotList
Christian Urban <urbanc@in.tum.de>
parents: 829
diff changeset
    40
  unfolding reflp_def symp_def transp_def
934
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    41
  apply(simp_all add: split_sum_all)
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    42
  apply(blast intro: equivp_reflp[OF a] equivp_reflp[OF b])
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    43
  apply(blast intro: equivp_symp[OF a] equivp_symp[OF b])
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    44
  apply(blast intro: equivp_transp[OF a] equivp_transp[OF b])
924
5455b19ef138 re-inserted lemma in QuotList
Christian Urban <urbanc@in.tum.de>
parents: 829
diff changeset
    45
  done
0
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
698
ed44eaaea63e Option and Sum quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 613
diff changeset
    47
lemma sum_quotient[quot_thm]:
ed44eaaea63e Option and Sum quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 613
diff changeset
    48
  assumes q1: "Quotient R1 Abs1 Rep1"
ed44eaaea63e Option and Sum quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 613
diff changeset
    49
  assumes q2: "Quotient R2 Abs2 Rep2"
ed44eaaea63e Option and Sum quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 613
diff changeset
    50
  shows "Quotient (sum_rel R1 R2) (sum_map Abs1 Abs2) (sum_map Rep1 Rep2)"
ed44eaaea63e Option and Sum quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 613
diff changeset
    51
  unfolding Quotient_def
934
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    52
  apply(simp add: split_sum_all)
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    53
  apply(simp_all add: Quotient_abs_rep[OF q1] Quotient_rel_rep[OF q1])
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    54
  apply(simp_all add: Quotient_abs_rep[OF q2] Quotient_rel_rep[OF q2])
924
5455b19ef138 re-inserted lemma in QuotList
Christian Urban <urbanc@in.tum.de>
parents: 829
diff changeset
    55
  using q1 q2
5455b19ef138 re-inserted lemma in QuotList
Christian Urban <urbanc@in.tum.de>
parents: 829
diff changeset
    56
  unfolding Quotient_def
5455b19ef138 re-inserted lemma in QuotList
Christian Urban <urbanc@in.tum.de>
parents: 829
diff changeset
    57
  apply(blast)+
698
ed44eaaea63e Option and Sum quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 613
diff changeset
    58
  done
ed44eaaea63e Option and Sum quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 613
diff changeset
    59
934
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    60
lemma sum_Inl_rsp[quot_respect]:
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    61
  assumes q1: "Quotient R1 Abs1 Rep1"
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    62
  assumes q2: "Quotient R2 Abs2 Rep2"
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    63
  shows "(R1 ===> sum_rel R1 R2) Inl Inl"
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    64
  by simp
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    65
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    66
lemma sum_Inr_rsp[quot_respect]:
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    67
  assumes q1: "Quotient R1 Abs1 Rep1"
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    68
  assumes q2: "Quotient R2 Abs2 Rep2"
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    69
  shows "(R2 ===> sum_rel R1 R2) Inr Inr"
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    70
  by simp
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    71
1122
d1eaed018e5d Changes from Makarius's code review + some noticed fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 937
diff changeset
    72
lemma sum_Inl_prs[quot_preserve]:
934
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    73
  assumes q1: "Quotient R1 Abs1 Rep1"
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    74
  assumes q2: "Quotient R2 Abs2 Rep2"
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    75
  shows "(Rep1 ---> sum_map Abs1 Abs2) Inl = Inl"
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    76
  apply(simp add: expand_fun_eq)
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    77
  apply(simp add: Quotient_abs_rep[OF q1])
829
42b90994ac77 map and rel simps for all quotients; needed when changing the relations to aggregate ones.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 779
diff changeset
    78
  done
42b90994ac77 map and rel simps for all quotients; needed when changing the relations to aggregate ones.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 779
diff changeset
    79
1122
d1eaed018e5d Changes from Makarius's code review + some noticed fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 937
diff changeset
    80
lemma sum_Inr_prs[quot_preserve]:
934
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    81
  assumes q1: "Quotient R1 Abs1 Rep1"
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    82
  assumes q2: "Quotient R2 Abs2 Rep2"
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    83
  shows "(Rep2 ---> sum_map Abs1 Abs2) Inr = Inr"
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    84
  apply(simp add: expand_fun_eq)
0b15b83ded4a added prs and rsp lemmas for Inl and Inr
Christian Urban <urbanc@in.tum.de>
parents: 924
diff changeset
    85
  apply(simp add: Quotient_abs_rep[OF q2])
829
42b90994ac77 map and rel simps for all quotients; needed when changing the relations to aggregate ones.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 779
diff changeset
    86
  done
42b90994ac77 map and rel simps for all quotients; needed when changing the relations to aggregate ones.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 779
diff changeset
    87
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1122
diff changeset
    88
lemma sum_map_id[id_simps]:
936
Christian Urban <urbanc@in.tum.de>
parents: 934 929
diff changeset
    89
  shows "sum_map id id = id"
Christian Urban <urbanc@in.tum.de>
parents: 934 929
diff changeset
    90
  by (simp add: expand_fun_eq split_sum_all)
829
42b90994ac77 map and rel simps for all quotients; needed when changing the relations to aggregate ones.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 779
diff changeset
    91
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1122
diff changeset
    92
lemma sum_rel_eq[id_simps]:
936
Christian Urban <urbanc@in.tum.de>
parents: 934 929
diff changeset
    93
  shows "sum_rel (op =) (op =) = (op =)"
Christian Urban <urbanc@in.tum.de>
parents: 934 929
diff changeset
    94
  by (simp add: expand_fun_eq split_sum_all)
829
42b90994ac77 map and rel simps for all quotients; needed when changing the relations to aggregate ones.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 779
diff changeset
    95
42b90994ac77 map and rel simps for all quotients; needed when changing the relations to aggregate ones.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 779
diff changeset
    96
end