698
|
1 |
theory QuotOption
|
|
2 |
imports QuotMain
|
|
3 |
begin
|
|
4 |
|
|
5 |
fun
|
|
6 |
option_rel
|
|
7 |
where
|
|
8 |
"option_rel R None None = True"
|
|
9 |
| "option_rel R (Some x) None = False"
|
|
10 |
| "option_rel R None (Some x) = False"
|
|
11 |
| "option_rel R (Some x) (Some y) = R x y"
|
|
12 |
|
|
13 |
fun
|
|
14 |
option_map
|
|
15 |
where
|
|
16 |
"option_map f None = None"
|
|
17 |
| "option_map f (Some x) = Some (f x)"
|
|
18 |
|
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
19 |
declare [[map option = (option_map, option_rel)]]
|
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
20 |
|
698
|
21 |
|
|
22 |
lemma option_quotient[quot_thm]:
|
|
23 |
assumes q: "Quotient R Abs Rep"
|
|
24 |
shows "Quotient (option_rel R) (option_map Abs) (option_map Rep)"
|
923
|
25 |
apply(unfold Quotient_def)
|
|
26 |
apply(auto)
|
|
27 |
apply(case_tac a, simp_all add: Quotient_abs_rep[OF q] Quotient_rel_rep[OF q])
|
|
28 |
apply(case_tac a, simp_all add: Quotient_abs_rep[OF q] Quotient_rel_rep[OF q])
|
|
29 |
apply(case_tac [!] r)
|
|
30 |
apply(case_tac [!] s)
|
|
31 |
apply(simp_all add: Quotient_abs_rep[OF q] Quotient_rel_rep[OF q] )
|
698
|
32 |
using q
|
|
33 |
unfolding Quotient_def
|
923
|
34 |
apply(blast)+
|
|
35 |
done
|
|
36 |
|
698
|
37 |
lemma option_equivp[quot_equiv]:
|
|
38 |
assumes a: "equivp R"
|
|
39 |
shows "equivp (option_rel R)"
|
923
|
40 |
apply(rule equivpI)
|
|
41 |
unfolding reflp_def symp_def transp_def
|
|
42 |
apply(auto)
|
|
43 |
apply(case_tac [!] x)
|
|
44 |
apply(simp_all add: equivp_reflp[OF a])
|
|
45 |
apply(case_tac [!] y)
|
|
46 |
apply(simp_all add: equivp_symp[OF a])
|
|
47 |
apply(case_tac [!] z)
|
698
|
48 |
apply(simp_all)
|
923
|
49 |
apply(clarify)
|
698
|
50 |
apply(rule equivp_transp[OF a])
|
|
51 |
apply(assumption)+
|
|
52 |
done
|
|
53 |
|
923
|
54 |
lemma option_None_rsp[quot_respect]:
|
|
55 |
assumes q: "Quotient R Abs Rep"
|
|
56 |
shows "option_rel R None None"
|
|
57 |
by simp
|
|
58 |
|
|
59 |
lemma option_Some_rsp[quot_respect]:
|
|
60 |
assumes q: "Quotient R Abs Rep"
|
|
61 |
shows "(R ===> option_rel R) Some Some"
|
|
62 |
by simp
|
|
63 |
|
|
64 |
lemma option_None_prs[quot_preserve]:
|
|
65 |
assumes q: "Quotient R Abs Rep"
|
|
66 |
shows "option_map Abs None = None"
|
|
67 |
by simp
|
|
68 |
|
|
69 |
lemma option_Some_prs[quot_respect]:
|
|
70 |
assumes q: "Quotient R Abs Rep"
|
|
71 |
shows "(Rep ---> option_map Abs) Some = Some"
|
|
72 |
apply(simp add: expand_fun_eq)
|
|
73 |
apply(simp add: Quotient_abs_rep[OF q])
|
|
74 |
done
|
|
75 |
|
|
76 |
lemma option_map_id[id_simps]:
|
|
77 |
shows "option_map id \<equiv> id"
|
829
42b90994ac77
map and rel simps for all quotients; needed when changing the relations to aggregate ones.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
78 |
apply (rule eq_reflection)
|
923
|
79 |
apply (auto simp add: expand_fun_eq)
|
829
42b90994ac77
map and rel simps for all quotients; needed when changing the relations to aggregate ones.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
80 |
apply (case_tac x)
|
42b90994ac77
map and rel simps for all quotients; needed when changing the relations to aggregate ones.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
81 |
apply (auto)
|
42b90994ac77
map and rel simps for all quotients; needed when changing the relations to aggregate ones.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
82 |
done
|
42b90994ac77
map and rel simps for all quotients; needed when changing the relations to aggregate ones.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
83 |
|
923
|
84 |
lemma option_rel_eq[id_simps]:
|
925
|
85 |
shows "option_rel (op =) \<equiv> (op =)"
|
923
|
86 |
apply(rule eq_reflection)
|
|
87 |
apply(auto simp add: expand_fun_eq)
|
|
88 |
apply(case_tac x)
|
|
89 |
apply(case_tac [!] xa)
|
|
90 |
apply(auto)
|
829
42b90994ac77
map and rel simps for all quotients; needed when changing the relations to aggregate ones.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
91 |
done
|
42b90994ac77
map and rel simps for all quotients; needed when changing the relations to aggregate ones.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
92 |
|
42b90994ac77
map and rel simps for all quotients; needed when changing the relations to aggregate ones.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
93 |
end
|