| author | nipkow | 
| Mon, 27 Apr 2015 15:02:51 +0200 | |
| changeset 60148 | f0fc2378a479 | 
| parent 53015 | a1119cf551e8 | 
| child 80914 | d97fdabd9e2b | 
| 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  | 
||
| 45212 | 6  | 
type_synonym penv = "(pname \<times> com) list"  | 
| 43158 | 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)" |  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
13  | 
Seq: "\<lbrakk> pe \<turnstile> (c\<^sub>1,s\<^sub>1) \<Rightarrow> s\<^sub>2; pe \<turnstile> (c\<^sub>2,s\<^sub>2) \<Rightarrow> s\<^sub>3 \<rbrakk> \<Longrightarrow>  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
14  | 
pe \<turnstile> (c\<^sub>1;;c\<^sub>2, s\<^sub>1) \<Rightarrow> s\<^sub>3" |  | 
| 43158 | 15  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
16  | 
IfTrue: "\<lbrakk> bval b s; pe \<turnstile> (c\<^sub>1,s) \<Rightarrow> t \<rbrakk> \<Longrightarrow>  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
17  | 
pe \<turnstile> (IF b THEN c\<^sub>1 ELSE c\<^sub>2, s) \<Rightarrow> t" |  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
18  | 
IfFalse: "\<lbrakk> \<not>bval b s; pe \<turnstile> (c\<^sub>2,s) \<Rightarrow> t \<rbrakk> \<Longrightarrow>  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
19  | 
pe \<turnstile> (IF b THEN c\<^sub>1 ELSE c\<^sub>2, s) \<Rightarrow> t" |  | 
| 43158 | 20  | 
|
21  | 
WhileFalse: "\<not>bval b s \<Longrightarrow> pe \<turnstile> (WHILE b DO c,s) \<Rightarrow> s" |  | 
|
22  | 
WhileTrue:  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
23  | 
"\<lbrakk> bval b s\<^sub>1; pe \<turnstile> (c,s\<^sub>1) \<Rightarrow> s\<^sub>2; pe \<turnstile> (WHILE b DO c, s\<^sub>2) \<Rightarrow> s\<^sub>3 \<rbrakk> \<Longrightarrow>  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
24  | 
pe \<turnstile> (WHILE b DO c, s\<^sub>1) \<Rightarrow> s\<^sub>3" |  | 
| 43158 | 25  | 
|
| 
52046
 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 
nipkow 
parents: 
51019 
diff
changeset
 | 
26  | 
Var: "pe \<turnstile> (c,s) \<Rightarrow> t  \<Longrightarrow>  pe \<turnstile> ({VAR x; c}, s) \<Rightarrow> t(x := s x)" |
 | 
| 43158 | 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  | 
||
| 
52046
 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 
nipkow 
parents: 
51019 
diff
changeset
 | 
32  | 
Proc: "(p,cp)#pe \<turnstile> (c,s) \<Rightarrow> t  \<Longrightarrow>  pe \<turnstile> ({PROC p = cp; c}, s) \<Rightarrow> t"
 | 
| 43158 | 33  | 
|
34  | 
code_pred big_step .  | 
|
35  | 
||
| 51019 | 36  | 
values "{map t [''x'', ''y''] |t. [] \<turnstile> (test_com, <>) \<Rightarrow> t}"
 | 
| 43158 | 37  | 
|
38  | 
end  |