author | nipkow |
Thu, 27 Jun 2013 10:11:11 +0200 | |
changeset 52460 | 92ae850a9bfd |
parent 51040 | faf7f0d4f9eb |
child 53015 | a1119cf551e8 |
permissions | -rw-r--r-- |
43141 | 1 |
header "Arithmetic and Boolean Expressions" |
2 |
||
3 |
theory AExp imports Main begin |
|
4 |
||
5 |
subsection "Arithmetic Expressions" |
|
6 |
||
45212 | 7 |
type_synonym vname = string |
43141 | 8 |
type_synonym val = int |
45212 | 9 |
type_synonym state = "vname \<Rightarrow> val" |
43141 | 10 |
|
49191 | 11 |
text_raw{*\snip{AExpaexpdef}{2}{1}{% *} |
45212 | 12 |
datatype aexp = N int | V vname | Plus aexp aexp |
49191 | 13 |
text_raw{*}%endsnip*} |
43141 | 14 |
|
49191 | 15 |
text_raw{*\snip{AExpavaldef}{1}{2}{% *} |
43141 | 16 |
fun aval :: "aexp \<Rightarrow> state \<Rightarrow> val" where |
45216 | 17 |
"aval (N n) s = n" | |
43141 | 18 |
"aval (V x) s = s x" | |
52460 | 19 |
"aval (Plus a\<^isub>1 a\<^isub>2) s = aval a\<^isub>1 s + aval a\<^isub>2 s" |
49191 | 20 |
text_raw{*}%endsnip*} |
43141 | 21 |
|
22 |
||
44923 | 23 |
value "aval (Plus (V ''x'') (N 5)) (\<lambda>x. if x = ''x'' then 7 else 0)" |
43141 | 24 |
|
43158 | 25 |
text {* The same state more concisely: *} |
44923 | 26 |
value "aval (Plus (V ''x'') (N 5)) ((\<lambda>x. 0) (''x'':= 7))" |
43158 | 27 |
|
28 |
text {* A little syntax magic to write larger states compactly: *} |
|
29 |
||
44923 | 30 |
definition null_state ("<>") where |
31 |
"null_state \<equiv> \<lambda>x. 0" |
|
44036 | 32 |
syntax |
33 |
"_State" :: "updbinds => 'a" ("<_>") |
|
43158 | 34 |
translations |
51040 | 35 |
"_State ms" == "_Update <> ms" |
43141 | 36 |
|
49191 | 37 |
text {* \noindent |
43158 | 38 |
We can now write a series of updates to the function @{text "\<lambda>x. 0"} compactly: |
39 |
*} |
|
44923 | 40 |
lemma "<a := Suc 0, b := 2> = (<> (a := Suc 0)) (b := 2)" |
43158 | 41 |
by (rule refl) |
42 |
||
44036 | 43 |
value "aval (Plus (V ''x'') (N 5)) <''x'' := 7>" |
43158 | 44 |
|
44923 | 45 |
|
49191 | 46 |
text {* In the @{term[source] "<a := b>"} syntax, variables that are not mentioned are 0 by default: |
44923 | 47 |
*} |
44036 | 48 |
value "aval (Plus (V ''x'') (N 5)) <''y'' := 7>" |
43141 | 49 |
|
44923 | 50 |
text{* Note that this @{text"<\<dots>>"} syntax works for any function space |
51 |
@{text"\<tau>\<^isub>1 \<Rightarrow> \<tau>\<^isub>2"} where @{text "\<tau>\<^isub>2"} has a @{text 0}. *} |
|
52 |
||
43141 | 53 |
|
45255 | 54 |
subsection "Constant Folding" |
43141 | 55 |
|
56 |
text{* Evaluate constant subsexpressions: *} |
|
57 |
||
49191 | 58 |
text_raw{*\snip{AExpasimpconstdef}{0}{2}{% *} |
43141 | 59 |
fun asimp_const :: "aexp \<Rightarrow> aexp" where |
60 |
"asimp_const (N n) = N n" | |
|
61 |
"asimp_const (V x) = V x" | |
|
45246
4fbeabee6487
added isaverbatimwrite that allows to cut out snippets of thy files in their latex form and dump them in a file
nipkow
parents:
45238
diff
changeset
|
62 |
"asimp_const (Plus a\<^isub>1 a\<^isub>2) = |
4fbeabee6487
added isaverbatimwrite that allows to cut out snippets of thy files in their latex form and dump them in a file
nipkow
parents:
45238
diff
changeset
|
63 |
(case (asimp_const a\<^isub>1, asimp_const a\<^isub>2) of |
4fbeabee6487
added isaverbatimwrite that allows to cut out snippets of thy files in their latex form and dump them in a file
nipkow
parents:
45238
diff
changeset
|
64 |
(N n\<^isub>1, N n\<^isub>2) \<Rightarrow> N(n\<^isub>1+n\<^isub>2) | |
4fbeabee6487
added isaverbatimwrite that allows to cut out snippets of thy files in their latex form and dump them in a file
nipkow
parents:
45238
diff
changeset
|
65 |
(b\<^isub>1,b\<^isub>2) \<Rightarrow> Plus b\<^isub>1 b\<^isub>2)" |
49191 | 66 |
text_raw{*}%endsnip*} |
43141 | 67 |
|
45238 | 68 |
theorem aval_asimp_const: |
43141 | 69 |
"aval (asimp_const a) s = aval a s" |
45015 | 70 |
apply(induction a) |
43141 | 71 |
apply (auto split: aexp.split) |
72 |
done |
|
73 |
||
74 |
text{* Now we also eliminate all occurrences 0 in additions. The standard |
|
75 |
method: optimized versions of the constructors: *} |
|
76 |
||
49191 | 77 |
text_raw{*\snip{AExpplusdef}{0}{2}{% *} |
43141 | 78 |
fun plus :: "aexp \<Rightarrow> aexp \<Rightarrow> aexp" where |
45246
4fbeabee6487
added isaverbatimwrite that allows to cut out snippets of thy files in their latex form and dump them in a file
nipkow
parents:
45238
diff
changeset
|
79 |
"plus (N i\<^isub>1) (N i\<^isub>2) = N(i\<^isub>1+i\<^isub>2)" | |
43141 | 80 |
"plus (N i) a = (if i=0 then a else Plus (N i) a)" | |
81 |
"plus a (N i) = (if i=0 then a else Plus a (N i))" | |
|
45246
4fbeabee6487
added isaverbatimwrite that allows to cut out snippets of thy files in their latex form and dump them in a file
nipkow
parents:
45238
diff
changeset
|
82 |
"plus a\<^isub>1 a\<^isub>2 = Plus a\<^isub>1 a\<^isub>2" |
49191 | 83 |
text_raw{*}%endsnip*} |
43141 | 84 |
|
85 |
lemma aval_plus[simp]: |
|
86 |
"aval (plus a1 a2) s = aval a1 s + aval a2 s" |
|
45015 | 87 |
apply(induction a1 a2 rule: plus.induct) |
43141 | 88 |
apply simp_all (* just for a change from auto *) |
89 |
done |
|
90 |
||
49191 | 91 |
text_raw{*\snip{AExpasimpdef}{2}{0}{% *} |
43141 | 92 |
fun asimp :: "aexp \<Rightarrow> aexp" where |
93 |
"asimp (N n) = N n" | |
|
94 |
"asimp (V x) = V x" | |
|
45246
4fbeabee6487
added isaverbatimwrite that allows to cut out snippets of thy files in their latex form and dump them in a file
nipkow
parents:
45238
diff
changeset
|
95 |
"asimp (Plus a\<^isub>1 a\<^isub>2) = plus (asimp a\<^isub>1) (asimp a\<^isub>2)" |
49191 | 96 |
text_raw{*}%endsnip*} |
43141 | 97 |
|
98 |
text{* Note that in @{const asimp_const} the optimized constructor was |
|
99 |
inlined. Making it a separate function @{const plus} improves modularity of |
|
100 |
the code and the proofs. *} |
|
101 |
||
102 |
value "asimp (Plus (Plus (N 0) (N 0)) (Plus (V ''x'') (N 0)))" |
|
103 |
||
104 |
theorem aval_asimp[simp]: |
|
105 |
"aval (asimp a) s = aval a s" |
|
45015 | 106 |
apply(induction a) |
43141 | 107 |
apply simp_all |
108 |
done |
|
109 |
||
110 |
end |