src/HOL/IMP/Vars.thy
author nipkow
Sat, 20 Apr 2013 20:57:49 +0200
changeset 51722 3da99469cc1b
parent 51698 c0af8bbc5825
child 52046 bc01725d7918
permissions -rw-r--r--
proved termination for fun-based AI
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
50303
nipkow
parents: 50173
diff changeset
     3
header "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
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    10
text{* We need to collect the variables in both arithmetic and boolean
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    11
expressions. For a change we do not introduce two functions, e.g.\ @{text
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    12
avars} and @{text bvars}, but we overload the name @{text vars}
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    13
via a \emph{type class}, a device that originated with Haskell: *}
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    14
 
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    15
class vars =
45212
e87feee00a4c renamed name -> vname
nipkow
parents: 45200
diff changeset
    16
fixes vars :: "'a \<Rightarrow> vname set"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    17
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    18
text{* This defines a type class ``vars'' with a single
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    19
function of (coincidentally) the same name. Then we define two separated
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    20
instances of the class, one for @{typ aexp} and one for @{typ bexp}: *}
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    21
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    22
instantiation aexp :: vars
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    23
begin
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    24
45212
e87feee00a4c renamed name -> vname
nipkow
parents: 45200
diff changeset
    25
fun vars_aexp :: "aexp \<Rightarrow> vname set" where
45774
nipkow
parents: 45716
diff changeset
    26
"vars (N n) = {}" |
nipkow
parents: 45716
diff changeset
    27
"vars (V x) = {x}" |
nipkow
parents: 45716
diff changeset
    28
"vars (Plus a\<^isub>1 a\<^isub>2) = vars a\<^isub>1 \<union> vars a\<^isub>2"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    29
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    30
instance ..
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    31
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    32
end
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    33
50173
nipkow
parents: 50006
diff changeset
    34
value "vars (Plus (V ''x'') (V ''y''))"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    35
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    36
instantiation bexp :: vars
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    37
begin
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    38
45212
e87feee00a4c renamed name -> vname
nipkow
parents: 45200
diff changeset
    39
fun vars_bexp :: "bexp \<Rightarrow> vname set" where
45774
nipkow
parents: 45716
diff changeset
    40
"vars (Bc v) = {}" |
nipkow
parents: 45716
diff changeset
    41
"vars (Not b) = vars b" |
nipkow
parents: 45716
diff changeset
    42
"vars (And b\<^isub>1 b\<^isub>2) = vars b\<^isub>1 \<union> vars b\<^isub>2" |
nipkow
parents: 45716
diff changeset
    43
"vars (Less a\<^isub>1 a\<^isub>2) = vars a\<^isub>1 \<union> vars a\<^isub>2"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    44
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    45
instance ..
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    46
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    47
end
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    48
50173
nipkow
parents: 50006
diff changeset
    49
value "vars (Less (Plus (V ''z'') (V ''y'')) (V ''x''))"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    50
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    51
abbreviation
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    52
  eq_on :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> bool"
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    53
 ("(_ =/ _/ on _)" [50,0,50] 50) where
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    54
"f = g on X == \<forall> x \<in> X. f x = g x"
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    55
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    56
lemma aval_eq_if_eq_on_vars[simp]:
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    57
  "s\<^isub>1 = s\<^isub>2 on vars a \<Longrightarrow> aval a s\<^isub>1 = aval a s\<^isub>2"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 43158
diff changeset
    58
apply(induction a)
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    59
apply simp_all
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    60
done
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    61
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    62
lemma bval_eq_if_eq_on_vars:
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    63
  "s\<^isub>1 = s\<^isub>2 on vars b \<Longrightarrow> bval b s\<^isub>1 = bval b s\<^isub>2"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 43158
diff changeset
    64
proof(induction b)
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    65
  case (Less a1 a2)
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    66
  hence "aval a1 s\<^isub>1 = aval a1 s\<^isub>2" and "aval a2 s\<^isub>1 = aval a2 s\<^isub>2" by simp_all
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    67
  thus ?case by simp
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    68
qed simp_all
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    69
51698
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    70
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    71
instantiation com :: vars
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    72
begin
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    73
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    74
fun vars_com :: "com \<Rightarrow> vname set" where
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    75
"vars com.SKIP = {}" |
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    76
"vars (x::=e) = {x} \<union> vars e" |
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    77
"vars (c1;c2) = vars c1 \<union> vars c2" |
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    78
"vars (IF b THEN c1 ELSE c2) = vars b \<union> vars c1 \<union> vars c2" |
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    79
"vars (WHILE b DO c) = vars b \<union> vars c"
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    80
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    81
instance ..
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    82
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    83
end
51698
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    84
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    85
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    86
lemma finite_avars[simp]: "finite(vars(a::aexp))"
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    87
by(induction a) simp_all
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    88
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    89
lemma finite_bvars[simp]: "finite(vars(b::bexp))"
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    90
by(induction b) (simp_all add: finite_avars)
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    91
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    92
lemma finite_cvars[simp]: "finite(vars(c::com))"
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    93
by(induction c) (simp_all add: finite_avars finite_bvars)
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    94
c0af8bbc5825 reduced duplication
nipkow
parents: 50303
diff changeset
    95
end