author | wenzelm |
Fri, 20 Sep 2024 19:51:08 +0200 | |
changeset 80914 | d97fdabd9e2b |
parent 69597 | ff784d5a5bfb |
permissions | -rw-r--r-- |
43158 | 1 |
(* Author: Tobias Nipkow *) |
2 |
||
58889 | 3 |
section "Definite Initialization Analysis" |
43158 | 4 |
|
51698 | 5 |
theory Vars imports Com |
43158 | 6 |
begin |
7 |
||
8 |
subsection "The Variables in an Expression" |
|
9 |
||
67406 | 10 |
text\<open>We need to collect the variables in both arithmetic and boolean |
69505 | 11 |
expressions. For a change we do not introduce two functions, e.g.\ \<open>avars\<close> and \<open>bvars\<close>, but we overload the name \<open>vars\<close> |
67406 | 12 |
via a \emph{type class}, a device that originated with Haskell:\<close> |
43158 | 13 |
|
14 |
class vars = |
|
45212 | 15 |
fixes vars :: "'a \<Rightarrow> vname set" |
43158 | 16 |
|
67406 | 17 |
text\<open>This defines a type class ``vars'' with a single |
43158 | 18 |
function of (coincidentally) the same name. Then we define two separated |
69597 | 19 |
instances of the class, one for \<^typ>\<open>aexp\<close> and one for \<^typ>\<open>bexp\<close>:\<close> |
43158 | 20 |
|
21 |
instantiation aexp :: vars |
|
22 |
begin |
|
23 |
||
45212 | 24 |
fun vars_aexp :: "aexp \<Rightarrow> vname set" where |
45774 | 25 |
"vars (N n) = {}" | |
26 |
"vars (V x) = {x}" | |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52072
diff
changeset
|
27 |
"vars (Plus a\<^sub>1 a\<^sub>2) = vars a\<^sub>1 \<union> vars a\<^sub>2" |
43158 | 28 |
|
29 |
instance .. |
|
30 |
||
31 |
end |
|
32 |
||
50173 | 33 |
value "vars (Plus (V ''x'') (V ''y''))" |
43158 | 34 |
|
35 |
instantiation bexp :: vars |
|
36 |
begin |
|
37 |
||
45212 | 38 |
fun vars_bexp :: "bexp \<Rightarrow> vname set" where |
45774 | 39 |
"vars (Bc v) = {}" | |
40 |
"vars (Not b) = vars b" | |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52072
diff
changeset
|
41 |
"vars (And b\<^sub>1 b\<^sub>2) = vars b\<^sub>1 \<union> vars b\<^sub>2" | |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52072
diff
changeset
|
42 |
"vars (Less a\<^sub>1 a\<^sub>2) = vars a\<^sub>1 \<union> vars a\<^sub>2" |
43158 | 43 |
|
44 |
instance .. |
|
45 |
||
46 |
end |
|
47 |
||
50173 | 48 |
value "vars (Less (Plus (V ''z'') (V ''y'')) (V ''x''))" |
43158 | 49 |
|
50 |
abbreviation |
|
51 |
eq_on :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> bool" |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
69597
diff
changeset
|
52 |
(\<open>(_ =/ _/ on _)\<close> [50,0,50] 50) where |
43158 | 53 |
"f = g on X == \<forall> x \<in> X. f x = g x" |
54 |
||
55 |
lemma aval_eq_if_eq_on_vars[simp]: |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52072
diff
changeset
|
56 |
"s\<^sub>1 = s\<^sub>2 on vars a \<Longrightarrow> aval a s\<^sub>1 = aval a s\<^sub>2" |
45015 | 57 |
apply(induction a) |
43158 | 58 |
apply simp_all |
59 |
done |
|
60 |
||
61 |
lemma bval_eq_if_eq_on_vars: |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52072
diff
changeset
|
62 |
"s\<^sub>1 = s\<^sub>2 on vars b \<Longrightarrow> bval b s\<^sub>1 = bval b s\<^sub>2" |
45015 | 63 |
proof(induction b) |
43158 | 64 |
case (Less a1 a2) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52072
diff
changeset
|
65 |
hence "aval a1 s\<^sub>1 = aval a1 s\<^sub>2" and "aval a2 s\<^sub>1 = aval a2 s\<^sub>2" by simp_all |
43158 | 66 |
thus ?case by simp |
67 |
qed simp_all |
|
68 |
||
52072 | 69 |
fun lvars :: "com \<Rightarrow> vname set" where |
70 |
"lvars SKIP = {}" | |
|
71 |
"lvars (x::=e) = {x}" | |
|
72 |
"lvars (c1;;c2) = lvars c1 \<union> lvars c2" | |
|
73 |
"lvars (IF b THEN c1 ELSE c2) = lvars c1 \<union> lvars c2" | |
|
74 |
"lvars (WHILE b DO c) = lvars c" |
|
75 |
||
76 |
fun rvars :: "com \<Rightarrow> vname set" where |
|
77 |
"rvars SKIP = {}" | |
|
78 |
"rvars (x::=e) = vars e" | |
|
79 |
"rvars (c1;;c2) = rvars c1 \<union> rvars c2" | |
|
80 |
"rvars (IF b THEN c1 ELSE c2) = vars b \<union> rvars c1 \<union> rvars c2" | |
|
81 |
"rvars (WHILE b DO c) = vars b \<union> rvars c" |
|
51698 | 82 |
|
83 |
instantiation com :: vars |
|
84 |
begin |
|
85 |
||
52072 | 86 |
definition "vars_com c = lvars c \<union> rvars c" |
51698 | 87 |
|
88 |
instance .. |
|
89 |
||
43158 | 90 |
end |
51698 | 91 |
|
52072 | 92 |
lemma vars_com_simps[simp]: |
93 |
"vars SKIP = {}" |
|
94 |
"vars (x::=e) = {x} \<union> vars e" |
|
95 |
"vars (c1;;c2) = vars c1 \<union> vars c2" |
|
96 |
"vars (IF b THEN c1 ELSE c2) = vars b \<union> vars c1 \<union> vars c2" |
|
97 |
"vars (WHILE b DO c) = vars b \<union> vars c" |
|
98 |
by(auto simp: vars_com_def) |
|
51698 | 99 |
|
100 |
lemma finite_avars[simp]: "finite(vars(a::aexp))" |
|
101 |
by(induction a) simp_all |
|
102 |
||
103 |
lemma finite_bvars[simp]: "finite(vars(b::bexp))" |
|
52072 | 104 |
by(induction b) simp_all |
105 |
||
106 |
lemma finite_lvars[simp]: "finite(lvars(c))" |
|
107 |
by(induction c) simp_all |
|
108 |
||
109 |
lemma finite_rvars[simp]: "finite(rvars(c))" |
|
110 |
by(induction c) simp_all |
|
51698 | 111 |
|
112 |
lemma finite_cvars[simp]: "finite(vars(c::com))" |
|
52072 | 113 |
by(simp add: vars_com_def) |
51698 | 114 |
|
115 |
end |