author | wenzelm |
Thu, 29 Aug 2019 17:13:49 +0200 | |
changeset 70634 | 0f8742b5a9e8 |
parent 68778 | 4566bac4517d |
permissions | -rw-r--r-- |
47619 | 1 |
(* Author: Tobias Nipkow *) |
2 |
||
68778 | 3 |
subsection "Computable State" |
4 |
||
47619 | 5 |
theory Abs_State |
6 |
imports Abs_Int0 |
|
7 |
begin |
|
8 |
||
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
9 |
type_synonym 'a st_rep = "(vname * 'a) list" |
47619 | 10 |
|
54930 | 11 |
fun fun_rep :: "('a::top) st_rep \<Rightarrow> vname \<Rightarrow> 'a" where |
51715
17b992f19b51
avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents:
51711
diff
changeset
|
12 |
"fun_rep [] = (\<lambda>x. \<top>)" | |
17b992f19b51
avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents:
51711
diff
changeset
|
13 |
"fun_rep ((x,a)#ps) = (fun_rep ps) (x := a)" |
17b992f19b51
avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents:
51711
diff
changeset
|
14 |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67406
diff
changeset
|
15 |
lemma fun_rep_map_of[code]: \<comment> \<open>original def is too slow\<close> |
51715
17b992f19b51
avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents:
51711
diff
changeset
|
16 |
"fun_rep ps = (%x. case map_of ps x of None \<Rightarrow> \<top> | Some a \<Rightarrow> a)" |
17b992f19b51
avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents:
51711
diff
changeset
|
17 |
by(induction ps rule: fun_rep.induct) auto |
47619 | 18 |
|
54930 | 19 |
definition eq_st :: "('a::top) st_rep \<Rightarrow> 'a st_rep \<Rightarrow> bool" where |
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
20 |
"eq_st S1 S2 = (fun_rep S1 = fun_rep S2)" |
47619 | 21 |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67406
diff
changeset
|
22 |
hide_type st \<comment> \<open>hide previous def to avoid long names\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67406
diff
changeset
|
23 |
declare [[typedef_overloaded]] \<comment> \<open>allow quotient types to depend on classes\<close> |
51720 | 24 |
|
61430 | 25 |
quotient_type 'a st = "('a::top) st_rep" / eq_st |
51800 | 26 |
morphisms rep_st St |
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
27 |
by (metis eq_st_def equivpI reflpI sympI transpI) |
47619 | 28 |
|
54930 | 29 |
lift_definition update :: "('a::top) st \<Rightarrow> vname \<Rightarrow> 'a \<Rightarrow> 'a st" |
51715
17b992f19b51
avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents:
51711
diff
changeset
|
30 |
is "\<lambda>ps x a. (x,a)#ps" |
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
31 |
by(auto simp: eq_st_def) |
47619 | 32 |
|
54930 | 33 |
lift_definition "fun" :: "('a::top) st \<Rightarrow> vname \<Rightarrow> 'a" is fun_rep |
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
34 |
by(simp add: eq_st_def) |
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
35 |
|
54930 | 36 |
definition show_st :: "vname set \<Rightarrow> ('a::top) st \<Rightarrow> (vname * 'a)set" where |
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
37 |
"show_st X S = (\<lambda>x. (x, fun S x)) ` X" |
47619 | 38 |
|
55466 | 39 |
definition "show_acom C = map_acom (map_option (show_st (vars(strip C)))) C" |
40 |
definition "show_acom_opt = map_option show_acom" |
|
47619 | 41 |
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
42 |
lemma fun_update[simp]: "fun (update S x y) = (fun S)(x:=y)" |
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
43 |
by transfer auto |
47619 | 44 |
|
54930 | 45 |
definition \<gamma>_st :: "(('a::top) \<Rightarrow> 'b set) \<Rightarrow> 'a st \<Rightarrow> (vname \<Rightarrow> 'b) set" where |
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
46 |
"\<gamma>_st \<gamma> F = {f. \<forall>x. f x \<in> \<gamma>(fun F x)}" |
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
47 |
|
54930 | 48 |
instantiation st :: (order_top) order |
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
49 |
begin |
47619 | 50 |
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
51 |
definition less_eq_st_rep :: "'a st_rep \<Rightarrow> 'a st_rep \<Rightarrow> bool" where |
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
52 |
"less_eq_st_rep ps1 ps2 = |
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
53 |
((\<forall>x \<in> set(map fst ps1) \<union> set(map fst ps2). fun_rep ps1 x \<le> fun_rep ps2 x))" |
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
54 |
|
51715
17b992f19b51
avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents:
51711
diff
changeset
|
55 |
lemma less_eq_st_rep_iff: |
17b992f19b51
avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents:
51711
diff
changeset
|
56 |
"less_eq_st_rep r1 r2 = (\<forall>x. fun_rep r1 x \<le> fun_rep r2 x)" |
17b992f19b51
avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents:
51711
diff
changeset
|
57 |
apply(auto simp: less_eq_st_rep_def fun_rep_map_of split: option.split) |
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
58 |
apply (metis Un_iff map_of_eq_None_iff option.distinct(1)) |
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
59 |
apply (metis Un_iff map_of_eq_None_iff option.distinct(1)) |
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
60 |
done |
47619 | 61 |
|
51715
17b992f19b51
avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents:
51711
diff
changeset
|
62 |
corollary less_eq_st_rep_iff_fun: |
17b992f19b51
avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents:
51711
diff
changeset
|
63 |
"less_eq_st_rep r1 r2 = (fun_rep r1 \<le> fun_rep r2)" |
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
64 |
by (metis less_eq_st_rep_iff le_fun_def) |
47619 | 65 |
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
66 |
lift_definition less_eq_st :: "'a st \<Rightarrow> 'a st \<Rightarrow> bool" is less_eq_st_rep |
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
67 |
by(auto simp add: eq_st_def less_eq_st_rep_iff) |
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
68 |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
69 |
definition less_st where "F < (G::'a st) = (F \<le> G \<and> \<not> G \<le> F)" |
47619 | 70 |
|
71 |
instance |
|
61179 | 72 |
proof (standard, goal_cases) |
73 |
case 1 show ?case by(rule less_st_def) |
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
74 |
next |
61179 | 75 |
case 2 show ?case by transfer (auto simp: less_eq_st_rep_def) |
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
76 |
next |
61179 | 77 |
case 3 thus ?case by transfer (metis less_eq_st_rep_iff order_trans) |
78 |
next |
|
79 |
case 4 thus ?case |
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
80 |
by transfer (metis less_eq_st_rep_iff eq_st_def fun_eq_iff antisym) |
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
81 |
qed |
47619 | 82 |
|
83 |
end |
|
84 |
||
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
85 |
lemma le_st_iff: "(F \<le> G) = (\<forall>x. fun F x \<le> fun G x)" |
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
86 |
by transfer (rule less_eq_st_rep_iff) |
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
87 |
|
54930 | 88 |
fun map2_st_rep :: "('a::top \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'a st_rep \<Rightarrow> 'a st_rep \<Rightarrow> 'a st_rep" where |
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
89 |
"map2_st_rep f [] ps2 = map (%(x,y). (x, f \<top> y)) ps2" | |
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
90 |
"map2_st_rep f ((x,y)#ps1) ps2 = |
51715
17b992f19b51
avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents:
51711
diff
changeset
|
91 |
(let y2 = fun_rep ps2 x |
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
92 |
in (x,f y y2) # map2_st_rep f ps1 ps2)" |
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
93 |
|
51715
17b992f19b51
avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents:
51711
diff
changeset
|
94 |
lemma fun_rep_map2_rep[simp]: "f \<top> \<top> = \<top> \<Longrightarrow> |
17b992f19b51
avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents:
51711
diff
changeset
|
95 |
fun_rep (map2_st_rep f ps1 ps2) = (\<lambda>x. f (fun_rep ps1 x) (fun_rep ps2 x))" |
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
96 |
apply(induction f ps1 ps2 rule: map2_st_rep.induct) |
51715
17b992f19b51
avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents:
51711
diff
changeset
|
97 |
apply(simp add: fun_rep_map_of map_of_map fun_eq_iff split: option.split) |
17b992f19b51
avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents:
51711
diff
changeset
|
98 |
apply(fastforce simp: fun_rep_map_of fun_eq_iff split:option.splits) |
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
99 |
done |
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
100 |
|
51826 | 101 |
instantiation st :: (semilattice_sup_top) semilattice_sup_top |
47619 | 102 |
begin |
103 |
||
67399 | 104 |
lift_definition sup_st :: "'a st \<Rightarrow> 'a st \<Rightarrow> 'a st" is "map2_st_rep (\<squnion>)" |
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
105 |
by (simp add: eq_st_def) |
47619 | 106 |
|
51994 | 107 |
lift_definition top_st :: "'a st" is "[]" . |
47619 | 108 |
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
109 |
instance |
61179 | 110 |
proof (standard, goal_cases) |
111 |
case 1 show ?case by transfer (simp add:less_eq_st_rep_iff) |
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
112 |
next |
61179 | 113 |
case 2 show ?case by transfer (simp add:less_eq_st_rep_iff) |
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
114 |
next |
61179 | 115 |
case 3 thus ?case by transfer (simp add:less_eq_st_rep_iff) |
116 |
next |
|
117 |
case 4 show ?case by transfer (simp add:less_eq_st_rep_iff fun_rep_map_of) |
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
118 |
qed |
47619 | 119 |
|
120 |
end |
|
121 |
||
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
122 |
lemma fun_top: "fun \<top> = (\<lambda>x. \<top>)" |
51715
17b992f19b51
avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents:
51711
diff
changeset
|
123 |
by transfer simp |
47619 | 124 |
|
125 |
lemma mono_update[simp]: |
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
126 |
"a1 \<le> a2 \<Longrightarrow> S1 \<le> S2 \<Longrightarrow> update S1 x a1 \<le> update S2 x a2" |
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
127 |
by transfer (auto simp add: less_eq_st_rep_def) |
47619 | 128 |
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
129 |
lemma mono_fun: "S1 \<le> S2 \<Longrightarrow> fun S1 x \<le> fun S2 x" |
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
130 |
by transfer (simp add: less_eq_st_rep_iff) |
47619 | 131 |
|
52504 | 132 |
locale Gamma_semilattice = Val_semilattice where \<gamma>=\<gamma> |
51826 | 133 |
for \<gamma> :: "'av::semilattice_sup_top \<Rightarrow> val set" |
47619 | 134 |
begin |
135 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52729
diff
changeset
|
136 |
abbreviation \<gamma>\<^sub>s :: "'av st \<Rightarrow> state set" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52729
diff
changeset
|
137 |
where "\<gamma>\<^sub>s == \<gamma>_st \<gamma>" |
47619 | 138 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52729
diff
changeset
|
139 |
abbreviation \<gamma>\<^sub>o :: "'av st option \<Rightarrow> state set" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52729
diff
changeset
|
140 |
where "\<gamma>\<^sub>o == \<gamma>_option \<gamma>\<^sub>s" |
47619 | 141 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52729
diff
changeset
|
142 |
abbreviation \<gamma>\<^sub>c :: "'av st option acom \<Rightarrow> state set acom" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52729
diff
changeset
|
143 |
where "\<gamma>\<^sub>c == map_acom \<gamma>\<^sub>o" |
47619 | 144 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52729
diff
changeset
|
145 |
lemma gamma_s_top[simp]: "\<gamma>\<^sub>s \<top> = UNIV" |
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
146 |
by(auto simp: \<gamma>_st_def fun_top) |
47619 | 147 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52729
diff
changeset
|
148 |
lemma gamma_o_Top[simp]: "\<gamma>\<^sub>o \<top> = UNIV" |
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
149 |
by (simp add: top_option_def) |
47619 | 150 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52729
diff
changeset
|
151 |
lemma mono_gamma_s: "f \<le> g \<Longrightarrow> \<gamma>\<^sub>s f \<subseteq> \<gamma>\<^sub>s g" |
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51698
diff
changeset
|
152 |
by(simp add:\<gamma>_st_def le_st_iff subset_iff) (metis mono_gamma subsetD) |
47619 | 153 |
|
154 |
lemma mono_gamma_o: |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52729
diff
changeset
|
155 |
"S1 \<le> S2 \<Longrightarrow> \<gamma>\<^sub>o S1 \<subseteq> \<gamma>\<^sub>o S2" |
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
156 |
by(induction S1 S2 rule: less_eq_option.induct)(simp_all add: mono_gamma_s) |
47619 | 157 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52729
diff
changeset
|
158 |
lemma mono_gamma_c: "C1 \<le> C2 \<Longrightarrow> \<gamma>\<^sub>c C1 \<le> \<gamma>\<^sub>c C2" |
52019
a4cbca8f7342
finally: acom with pointwise access and update of annotations
nipkow
parents:
51826
diff
changeset
|
159 |
by (simp add: less_eq_acom_def mono_gamma_o size_annos anno_map_acom size_annos_same[of C1 C2]) |
47619 | 160 |
|
161 |
lemma in_gamma_option_iff: |
|
67613 | 162 |
"x \<in> \<gamma>_option r u \<longleftrightarrow> (\<exists>u'. u = Some u' \<and> x \<in> r u')" |
47619 | 163 |
by (cases u) auto |
164 |
||
165 |
end |
|
166 |
||
167 |
end |