| author | blanchet | 
| Tue, 15 Jul 2014 00:21:32 +0200 | |
| changeset 57554 | 12fb55fc11a6 | 
| 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: 
43158diff
changeset | 1 | theory Procs_Stat_Vars_Stat imports Procs | 
| 43158 | 2 | begin | 
| 3 | ||
| 4 | subsubsection "Static Scoping of Procedures and Variables" | |
| 5 | ||
| 6 | type_synonym addr = nat | |
| 45212 | 7 | type_synonym venv = "vname \<Rightarrow> addr" | 
| 43158 | 8 | type_synonym store = "addr \<Rightarrow> val" | 
| 45212 | 9 | type_synonym penv = "(pname \<times> com \<times> venv) list" | 
| 43158 | 10 | |
| 11 | fun venv :: "penv \<times> venv \<times> nat \<Rightarrow> venv" where | |
| 12 | "venv(_,ve,_) = ve" | |
| 13 | ||
| 14 | inductive | |
| 15 | big_step :: "penv \<times> venv \<times> nat \<Rightarrow> com \<times> store \<Rightarrow> store \<Rightarrow> bool" | |
| 16 |   ("_ \<turnstile> _ \<Rightarrow> _" [60,0,60] 55)
 | |
| 17 | where | |
| 18 | Skip: "e \<turnstile> (SKIP,s) \<Rightarrow> s" | | |
| 19 | Assign: "(pe,ve,f) \<turnstile> (x ::= a,s) \<Rightarrow> s(ve x := aval a (s o ve))" | | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52046diff
changeset | 20 | Seq: "\<lbrakk> e \<turnstile> (c\<^sub>1,s\<^sub>1) \<Rightarrow> s\<^sub>2; e \<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: 
52046diff
changeset | 21 | e \<turnstile> (c\<^sub>1;;c\<^sub>2, s\<^sub>1) \<Rightarrow> s\<^sub>3" | | 
| 43158 | 22 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52046diff
changeset | 23 | IfTrue: "\<lbrakk> bval b (s \<circ> venv e); e \<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: 
52046diff
changeset | 24 | e \<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: 
52046diff
changeset | 25 | IfFalse: "\<lbrakk> \<not>bval b (s \<circ> venv e); e \<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: 
52046diff
changeset | 26 | e \<turnstile> (IF b THEN c\<^sub>1 ELSE c\<^sub>2, s) \<Rightarrow> t" | | 
| 43158 | 27 | |
| 28 | WhileFalse: "\<not>bval b (s \<circ> venv e) \<Longrightarrow> e \<turnstile> (WHILE b DO c,s) \<Rightarrow> s" | | |
| 29 | WhileTrue: | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52046diff
changeset | 30 | "\<lbrakk> bval b (s\<^sub>1 \<circ> venv e); e \<turnstile> (c,s\<^sub>1) \<Rightarrow> s\<^sub>2; | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52046diff
changeset | 31 | e \<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: 
52046diff
changeset | 32 | e \<turnstile> (WHILE b DO c, s\<^sub>1) \<Rightarrow> s\<^sub>3" | | 
| 43158 | 33 | |
| 34 | Var: "(pe,ve(x:=f),f+1) \<turnstile> (c,s) \<Rightarrow> t \<Longrightarrow> | |
| 52046 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
51019diff
changeset | 35 |       (pe,ve,f) \<turnstile> ({VAR x; c}, s) \<Rightarrow> t" |
 | 
| 43158 | 36 | |
| 37 | Call1: "((p,c,ve)#pe,ve,f) \<turnstile> (c, s) \<Rightarrow> t \<Longrightarrow> | |
| 38 | ((p,c,ve)#pe,ve',f) \<turnstile> (CALL p, s) \<Rightarrow> t" | | |
| 39 | Call2: "\<lbrakk> p' \<noteq> p; (pe,ve,f) \<turnstile> (CALL p, s) \<Rightarrow> t \<rbrakk> \<Longrightarrow> | |
| 40 | ((p',c,ve')#pe,ve,f) \<turnstile> (CALL p, s) \<Rightarrow> t" | | |
| 41 | ||
| 42 | Proc: "((p,cp,ve)#pe,ve,f) \<turnstile> (c,s) \<Rightarrow> t | |
| 52046 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
51019diff
changeset | 43 |       \<Longrightarrow>  (pe,ve,f) \<turnstile> ({PROC p = cp; c}, s) \<Rightarrow> t"
 | 
| 43158 | 44 | |
| 45 | code_pred big_step . | |
| 46 | ||
| 44923 | 47 | |
| 51019 | 48 | values "{map t [10,11] |t.
 | 
| 49 | ([], <''x'' := 10, ''y'' := 11>, 12) | |
| 44923 | 50 | \<turnstile> (test_com, <>) \<Rightarrow> t}" | 
| 43158 | 51 | |
| 52 | end |