src/HOL/IMP/ASM.thy
author nipkow
Sun, 04 Nov 2018 12:07:24 +0100
changeset 69232 2b913054a9cf
parent 67406 23307fd33906
permissions -rw-r--r--
Faster Braun tree functions
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
58889
5b7a9633cfa8 modernized header uniformly as section;
wenzelm
parents: 58310
diff changeset
     1
section "Stack Machine and Compilation"
43141
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 ASM imports AExp begin
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
     4
45266
13b5fb92b9f5 tuned text
nipkow
parents: 45257
diff changeset
     5
subsection "Stack Machine"
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
     6
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67079
diff changeset
     7
text_raw\<open>\snip{ASMinstrdef}{0}{1}{%\<close>
58310
91ea607a34d8 updated news
blanchet
parents: 58249
diff changeset
     8
datatype instr = LOADI val | LOAD vname | ADD
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67079
diff changeset
     9
text_raw\<open>}%endsnip\<close>
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    10
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67079
diff changeset
    11
text_raw\<open>\snip{ASMstackdef}{1}{2}{%\<close>
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    12
type_synonym stack = "val list"
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67079
diff changeset
    13
text_raw\<open>}%endsnip\<close>
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    14
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67079
diff changeset
    15
text\<open>\noindent Abbreviations are transparent: they are unfolded after
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    16
parsing and folded back again before printing. Internally, they do not
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67079
diff changeset
    17
exist.\<close>
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    18
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67079
diff changeset
    19
text_raw\<open>\snip{ASMexeconedef}{0}{1}{%\<close>
45257
12063e071d92 renamed in ASM
nipkow
parents: 45222
diff changeset
    20
fun exec1 :: "instr \<Rightarrow> state \<Rightarrow> stack \<Rightarrow> stack" where
12063e071d92 renamed in ASM
nipkow
parents: 45222
diff changeset
    21
"exec1 (LOADI n) _ stk  =  n # stk" |
12063e071d92 renamed in ASM
nipkow
parents: 45222
diff changeset
    22
"exec1 (LOAD x) s stk  =  s(x) # stk" |
67079
nipkow
parents: 58889
diff changeset
    23
"exec1  ADD _ (j # i # stk)  =  (i + j) # stk"
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67079
diff changeset
    24
text_raw\<open>}%endsnip\<close>
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    25
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67079
diff changeset
    26
text_raw\<open>\snip{ASMexecdef}{1}{2}{%\<close>
45257
12063e071d92 renamed in ASM
nipkow
parents: 45222
diff changeset
    27
fun exec :: "instr list \<Rightarrow> state \<Rightarrow> stack \<Rightarrow> stack" where
12063e071d92 renamed in ASM
nipkow
parents: 45222
diff changeset
    28
"exec [] _ stk = stk" |
12063e071d92 renamed in ASM
nipkow
parents: 45222
diff changeset
    29
"exec (i#is) s stk = exec is s (exec1 i s stk)"
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67079
diff changeset
    30
text_raw\<open>}%endsnip\<close>
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    31
50007
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 49191
diff changeset
    32
value "exec [LOADI 5, LOAD ''y'', ADD] <''x'' := 42, ''y'' := 43> [50]"
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    33
45257
12063e071d92 renamed in ASM
nipkow
parents: 45222
diff changeset
    34
lemma exec_append[simp]:
12063e071d92 renamed in ASM
nipkow
parents: 45222
diff changeset
    35
  "exec (is1@is2) s stk = exec is2 s (exec is1 s stk)"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44036
diff changeset
    36
apply(induction is1 arbitrary: stk)
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    37
apply (auto)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    38
done
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    39
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    40
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    41
subsection "Compilation"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    42
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67079
diff changeset
    43
text_raw\<open>\snip{ASMcompdef}{0}{2}{%\<close>
45257
12063e071d92 renamed in ASM
nipkow
parents: 45222
diff changeset
    44
fun comp :: "aexp \<Rightarrow> instr list" where
12063e071d92 renamed in ASM
nipkow
parents: 45222
diff changeset
    45
"comp (N n) = [LOADI n]" |
12063e071d92 renamed in ASM
nipkow
parents: 45222
diff changeset
    46
"comp (V x) = [LOAD x]" |
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50007
diff changeset
    47
"comp (Plus e\<^sub>1 e\<^sub>2) = comp e\<^sub>1 @ comp e\<^sub>2 @ [ADD]"
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67079
diff changeset
    48
text_raw\<open>}%endsnip\<close>
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    49
45257
12063e071d92 renamed in ASM
nipkow
parents: 45222
diff changeset
    50
value "comp (Plus (Plus (V ''x'') (N 1)) (V ''z''))"
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    51
45257
12063e071d92 renamed in ASM
nipkow
parents: 45222
diff changeset
    52
theorem exec_comp: "exec (comp a) s stk = aval a s # stk"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44036
diff changeset
    53
apply(induction a arbitrary: stk)
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    54
apply (auto)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    55
done
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    56
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    57
end