src/HOL/IMP/Vars.thy
author wenzelm
Wed, 21 Jun 2023 11:05:20 +0200
changeset 78184 4309bcc8f28b
parent 69597 ff784d5a5bfb
child 80914 d97fdabd9e2b
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     2
58889
5b7a9633cfa8 modernized header uniformly as section;
wenzelm
parents: 53015
diff changeset
     3
section "Definite Initialization Analysis"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     4
51698
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
     5
theory Vars imports Com
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     6
begin
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     7
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     8
subsection "The Variables in an Expression"
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     9
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58889
diff changeset
    10
text\<open>We need to collect the variables in both arithmetic and boolean
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67406
diff changeset
    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
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58889
diff changeset
    12
via a \emph{type class}, a device that originated with Haskell:\<close>
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    13
 
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    14
class vars =
45212
e87feee00a4c renamed name -> vname
nipkow
parents: 45200
diff changeset
    15
fixes vars :: "'a \<Rightarrow> vname set"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    16
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58889
diff changeset
    17
text\<open>This defines a type class ``vars'' with a single
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    18
function of (coincidentally) the same name. Then we define two separated
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    19
instances of the class, one for \<^typ>\<open>aexp\<close> and one for \<^typ>\<open>bexp\<close>:\<close>
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    20
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    21
instantiation aexp :: vars
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    22
begin
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    23
45212
e87feee00a4c renamed name -> vname
nipkow
parents: 45200
diff changeset
    24
fun vars_aexp :: "aexp \<Rightarrow> vname set" where
45774
nipkow
parents: 45716
diff changeset
    25
"vars (N n) = {}" |
nipkow
parents: 45716
diff changeset
    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
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    28
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    29
instance ..
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    30
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    31
end
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    32
50173
nipkow
parents: 50006
diff changeset
    33
value "vars (Plus (V ''x'') (V ''y''))"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    34
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    35
instantiation bexp :: vars
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    36
begin
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    37
45212
e87feee00a4c renamed name -> vname
nipkow
parents: 45200
diff changeset
    38
fun vars_bexp :: "bexp \<Rightarrow> vname set" where
45774
nipkow
parents: 45716
diff changeset
    39
"vars (Bc v) = {}" |
nipkow
parents: 45716
diff changeset
    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
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    43
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    44
instance ..
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    45
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    46
end
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    47
50173
nipkow
parents: 50006
diff changeset
    48
value "vars (Less (Plus (V ''z'') (V ''y'')) (V ''x''))"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    49
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    50
abbreviation
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    51
  eq_on :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> bool"
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    52
 ("(_ =/ _/ on _)" [50,0,50] 50) where
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    53
"f = g on X == \<forall> x \<in> X. f x = g x"
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    54
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    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
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 43158
diff changeset
    57
apply(induction a)
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    58
apply simp_all
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    59
done
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    60
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    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
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 43158
diff changeset
    63
proof(induction b)
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    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
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    66
  thus ?case by simp
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    67
qed simp_all
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    68
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    69
fun lvars :: "com \<Rightarrow> vname set" where
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    70
"lvars SKIP = {}" |
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    71
"lvars (x::=e) = {x}" |
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    72
"lvars (c1;;c2) = lvars c1 \<union> lvars c2" |
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    73
"lvars (IF b THEN c1 ELSE c2) = lvars c1 \<union> lvars c2" |
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    74
"lvars (WHILE b DO c) = lvars c"
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    75
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    76
fun rvars :: "com \<Rightarrow> vname set" where
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    77
"rvars SKIP = {}" |
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    78
"rvars (x::=e) = vars e" |
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    79
"rvars (c1;;c2) = rvars c1 \<union> rvars c2" |
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    80
"rvars (IF b THEN c1 ELSE c2) = vars b \<union> rvars c1 \<union> rvars c2" |
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    81
"rvars (WHILE b DO c) = vars b \<union> rvars c"
51698
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    82
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    83
instantiation com :: vars
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    84
begin
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    85
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    86
definition "vars_com c = lvars c \<union> rvars c"
51698
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    87
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    88
instance ..
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    89
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    90
end
51698
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    91
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    92
lemma vars_com_simps[simp]:
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    93
  "vars SKIP = {}"
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    94
  "vars (x::=e) = {x} \<union> vars e"
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    95
  "vars (c1;;c2) = vars c1 \<union> vars c2"
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    96
  "vars (IF b THEN c1 ELSE c2) = vars b \<union> vars c1 \<union> vars c2"
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    97
  "vars (WHILE b DO c) = vars b \<union> vars c"
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    98
by(auto simp: vars_com_def)
51698
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    99
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
   100
lemma finite_avars[simp]: "finite(vars(a::aexp))"
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
   101
by(induction a) simp_all
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
   102
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
   103
lemma finite_bvars[simp]: "finite(vars(b::bexp))"
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   104
by(induction b) simp_all
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   105
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   106
lemma finite_lvars[simp]: "finite(lvars(c))"
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   107
by(induction c) simp_all
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   108
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   109
lemma finite_rvars[simp]: "finite(rvars(c))"
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   110
by(induction c) simp_all
51698
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
   111
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
   112
lemma finite_cvars[simp]: "finite(vars(c::com))"
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   113
by(simp add: vars_com_def)
51698
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
   114
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
   115
end