src/HOL/IMP/AExp.thy
author wenzelm
Wed, 17 Aug 2011 20:08:36 +0200
changeset 44246 380a4677c55d
parent 44036 d03f9f28d01d
child 44923 b80108b346a9
permissions -rw-r--r--
tuned signature;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
     1
header "Arithmetic and Boolean Expressions"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
     2
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
     3
theory AExp imports Main begin
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
     4
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
     5
subsection "Arithmetic Expressions"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
     6
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
     7
type_synonym name = string
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
     8
type_synonym val = int
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
     9
type_synonym state = "name \<Rightarrow> val"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    10
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    11
datatype aexp = N int | V name | Plus aexp aexp
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    12
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    13
fun aval :: "aexp \<Rightarrow> state \<Rightarrow> val" where
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    14
"aval (N n) _ = n" |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    15
"aval (V x) s = s x" |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    16
"aval (Plus a1 a2) s = aval a1 s + aval a2 s"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    17
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    18
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    19
value "aval (Plus (V ''x'') (N 5)) (%x. if x = ''x'' then 7 else 0)"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    20
43158
686fa0a0696e imported rest of new IMP
kleing
parents: 43141
diff changeset
    21
text {* The same state more concisely: *}
686fa0a0696e imported rest of new IMP
kleing
parents: 43141
diff changeset
    22
value "aval (Plus (V ''x'') (N 5)) ((%x. 0) (''x'':= 7))"
686fa0a0696e imported rest of new IMP
kleing
parents: 43141
diff changeset
    23
686fa0a0696e imported rest of new IMP
kleing
parents: 43141
diff changeset
    24
text {* A little syntax magic to write larger states compactly: *}
686fa0a0696e imported rest of new IMP
kleing
parents: 43141
diff changeset
    25
43250
c729110a9f08 use null_heap instead of %_. 0 to avoid printing problems
kleing
parents: 43158
diff changeset
    26
definition
c729110a9f08 use null_heap instead of %_. 0 to avoid printing problems
kleing
parents: 43158
diff changeset
    27
  "null_heap \<equiv> \<lambda>x. 0"
44036
d03f9f28d01d new state syntax with less conflicts
kleing
parents: 43250
diff changeset
    28
syntax 
d03f9f28d01d new state syntax with less conflicts
kleing
parents: 43250
diff changeset
    29
  "_State" :: "updbinds => 'a" ("<_>")
43158
686fa0a0696e imported rest of new IMP
kleing
parents: 43141
diff changeset
    30
translations
44036
d03f9f28d01d new state syntax with less conflicts
kleing
parents: 43250
diff changeset
    31
  "_State ms" => "_Update (CONST null_heap) ms"
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    32
43158
686fa0a0696e imported rest of new IMP
kleing
parents: 43141
diff changeset
    33
text {* 
686fa0a0696e imported rest of new IMP
kleing
parents: 43141
diff changeset
    34
  We can now write a series of updates to the function @{text "\<lambda>x. 0"} compactly:
686fa0a0696e imported rest of new IMP
kleing
parents: 43141
diff changeset
    35
*}
44036
d03f9f28d01d new state syntax with less conflicts
kleing
parents: 43250
diff changeset
    36
lemma "<a := Suc 0, b := 2> = (null_heap (a := Suc 0)) (b := 2)"
43158
686fa0a0696e imported rest of new IMP
kleing
parents: 43141
diff changeset
    37
  by (rule refl)
686fa0a0696e imported rest of new IMP
kleing
parents: 43141
diff changeset
    38
44036
d03f9f28d01d new state syntax with less conflicts
kleing
parents: 43250
diff changeset
    39
value "aval (Plus (V ''x'') (N 5)) <''x'' := 7>"
43158
686fa0a0696e imported rest of new IMP
kleing
parents: 43141
diff changeset
    40
686fa0a0696e imported rest of new IMP
kleing
parents: 43141
diff changeset
    41
text {* Variables that are not mentioned in the state are 0 by default in 
44036
d03f9f28d01d new state syntax with less conflicts
kleing
parents: 43250
diff changeset
    42
  the @{term "<a := b::int>"} syntax: 
43158
686fa0a0696e imported rest of new IMP
kleing
parents: 43141
diff changeset
    43
*}   
44036
d03f9f28d01d new state syntax with less conflicts
kleing
parents: 43250
diff changeset
    44
value "aval (Plus (V ''x'') (N 5)) <''y'' := 7>"
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    45
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    46
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    47
subsection "Optimization"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    48
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    49
text{* Evaluate constant subsexpressions: *}
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    50
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    51
fun asimp_const :: "aexp \<Rightarrow> aexp" where
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    52
"asimp_const (N n) = N n" |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    53
"asimp_const (V x) = V x" |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    54
"asimp_const (Plus a1 a2) =
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    55
  (case (asimp_const a1, asimp_const a2) of
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    56
    (N n1, N n2) \<Rightarrow> N(n1+n2) |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    57
    (a1',a2') \<Rightarrow> Plus a1' a2')"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    58
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    59
theorem aval_asimp_const[simp]:
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    60
  "aval (asimp_const a) s = aval a s"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    61
apply(induct a)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    62
apply (auto split: aexp.split)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    63
done
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    64
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    65
text{* Now we also eliminate all occurrences 0 in additions. The standard
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    66
method: optimized versions of the constructors: *}
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    67
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    68
fun plus :: "aexp \<Rightarrow> aexp \<Rightarrow> aexp" where
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    69
"plus (N i1) (N i2) = N(i1+i2)" |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    70
"plus (N i) a = (if i=0 then a else Plus (N i) a)" |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    71
"plus a (N i) = (if i=0 then a else Plus a (N i))" |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    72
"plus a1 a2 = Plus a1 a2"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    73
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    74
code_thms plus
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    75
code_thms plus
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    76
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    77
(* FIXME: dropping subsumed code eqns?? *)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    78
lemma aval_plus[simp]:
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    79
  "aval (plus a1 a2) s = aval a1 s + aval a2 s"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    80
apply(induct a1 a2 rule: plus.induct)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    81
apply simp_all (* just for a change from auto *)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    82
done
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    83
code_thms plus
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    84
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    85
fun asimp :: "aexp \<Rightarrow> aexp" where
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    86
"asimp (N n) = N n" |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    87
"asimp (V x) = V x" |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    88
"asimp (Plus a1 a2) = plus (asimp a1) (asimp a2)"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    89
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    90
text{* Note that in @{const asimp_const} the optimized constructor was
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    91
inlined. Making it a separate function @{const plus} improves modularity of
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    92
the code and the proofs. *}
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    93
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    94
value "asimp (Plus (Plus (N 0) (N 0)) (Plus (V ''x'') (N 0)))"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    95
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    96
theorem aval_asimp[simp]:
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    97
  "aval (asimp a) s = aval a s"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    98
apply(induct a)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    99
apply simp_all
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
   100
done
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
   101
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
   102
end