author | nipkow |
Wed, 14 Sep 2011 06:49:01 +0200 | |
changeset 44923 | b80108b346a9 |
parent 44177 | b4b5cbca2519 |
child 45212 | e87feee00a4c |
permissions | -rw-r--r-- |
44177
b4b5cbca2519
IMP/Util distinguishes between sets and functions again; imported only where used.
kleing
parents:
43158
diff
changeset
|
1 |
theory Procs_Stat_Vars_Dyn imports Procs |
43158 | 2 |
begin |
3 |
||
4 |
subsubsection "Static Scoping of Procedures, Dynamic of Variables" |
|
5 |
||
6 |
type_synonym penv = "(name \<times> com) list" |
|
7 |
||
8 |
inductive |
|
9 |
big_step :: "penv \<Rightarrow> com \<times> state \<Rightarrow> state \<Rightarrow> bool" ("_ \<turnstile> _ \<Rightarrow> _" [60,0,60] 55) |
|
10 |
where |
|
11 |
Skip: "pe \<turnstile> (SKIP,s) \<Rightarrow> s" | |
|
12 |
Assign: "pe \<turnstile> (x ::= a,s) \<Rightarrow> s(x := aval a s)" | |
|
13 |
Semi: "\<lbrakk> pe \<turnstile> (c\<^isub>1,s\<^isub>1) \<Rightarrow> s\<^isub>2; pe \<turnstile> (c\<^isub>2,s\<^isub>2) \<Rightarrow> s\<^isub>3 \<rbrakk> \<Longrightarrow> |
|
14 |
pe \<turnstile> (c\<^isub>1;c\<^isub>2, s\<^isub>1) \<Rightarrow> s\<^isub>3" | |
|
15 |
||
16 |
IfTrue: "\<lbrakk> bval b s; pe \<turnstile> (c\<^isub>1,s) \<Rightarrow> t \<rbrakk> \<Longrightarrow> |
|
17 |
pe \<turnstile> (IF b THEN c\<^isub>1 ELSE c\<^isub>2, s) \<Rightarrow> t" | |
|
18 |
IfFalse: "\<lbrakk> \<not>bval b s; pe \<turnstile> (c\<^isub>2,s) \<Rightarrow> t \<rbrakk> \<Longrightarrow> |
|
19 |
pe \<turnstile> (IF b THEN c\<^isub>1 ELSE c\<^isub>2, s) \<Rightarrow> t" | |
|
20 |
||
21 |
WhileFalse: "\<not>bval b s \<Longrightarrow> pe \<turnstile> (WHILE b DO c,s) \<Rightarrow> s" | |
|
22 |
WhileTrue: |
|
23 |
"\<lbrakk> bval b s\<^isub>1; pe \<turnstile> (c,s\<^isub>1) \<Rightarrow> s\<^isub>2; pe \<turnstile> (WHILE b DO c, s\<^isub>2) \<Rightarrow> s\<^isub>3 \<rbrakk> \<Longrightarrow> |
|
24 |
pe \<turnstile> (WHILE b DO c, s\<^isub>1) \<Rightarrow> s\<^isub>3" | |
|
25 |
||
26 |
Var: "pe \<turnstile> (c,s) \<Rightarrow> t \<Longrightarrow> pe \<turnstile> ({VAR x;; c}, s) \<Rightarrow> t(x := s x)" | |
|
27 |
||
28 |
Call1: "(p,c)#pe \<turnstile> (c, s) \<Rightarrow> t \<Longrightarrow> (p,c)#pe \<turnstile> (CALL p, s) \<Rightarrow> t" | |
|
29 |
Call2: "\<lbrakk> p' \<noteq> p; pe \<turnstile> (CALL p, s) \<Rightarrow> t \<rbrakk> \<Longrightarrow> |
|
30 |
(p',c)#pe \<turnstile> (CALL p, s) \<Rightarrow> t" | |
|
31 |
||
32 |
Proc: "(p,cp)#pe \<turnstile> (c,s) \<Rightarrow> t \<Longrightarrow> pe \<turnstile> ({PROC p = cp;; c}, s) \<Rightarrow> t" |
|
33 |
||
34 |
code_pred big_step . |
|
35 |
||
44923 | 36 |
values "{map t [''x'', ''y'', ''z''] |t. |
37 |
[] \<turnstile> (CALL ''p'', <''x'' := 42, ''y'' := 43>) \<Rightarrow> t}" |
|
43158 | 38 |
|
44923 | 39 |
values "{map t [''x'', ''y'', ''z''] |t. [] \<turnstile> (test_com, <>) \<Rightarrow> t}" |
43158 | 40 |
|
41 |
end |