src/HOL/IMP/Procs_Stat_Vars_Stat.thy
changeset 43158 686fa0a0696e
child 44177 b4b5cbca2519
equal deleted inserted replaced
43157:b505be6f029a 43158:686fa0a0696e
       
     1 theory Procs_Stat_Vars_Stat imports Util Procs
       
     2 begin
       
     3 
       
     4 subsubsection "Static Scoping of Procedures and Variables"
       
     5 
       
     6 type_synonym addr = nat
       
     7 type_synonym venv = "name \<Rightarrow> addr"
       
     8 type_synonym store = "addr \<Rightarrow> val"
       
     9 type_synonym penv = "(name \<times> com \<times> venv) list"
       
    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))" |
       
    20 Semi:    "\<lbrakk> e \<turnstile> (c\<^isub>1,s\<^isub>1) \<Rightarrow> s\<^isub>2;  e \<turnstile> (c\<^isub>2,s\<^isub>2) \<Rightarrow> s\<^isub>3 \<rbrakk> \<Longrightarrow>
       
    21           e \<turnstile> (c\<^isub>1;c\<^isub>2, s\<^isub>1) \<Rightarrow> s\<^isub>3" |
       
    22 
       
    23 IfTrue:  "\<lbrakk> bval b (s \<circ> venv e);  e \<turnstile> (c\<^isub>1,s) \<Rightarrow> t \<rbrakk> \<Longrightarrow>
       
    24          e \<turnstile> (IF b THEN c\<^isub>1 ELSE c\<^isub>2, s) \<Rightarrow> t" |
       
    25 IfFalse: "\<lbrakk> \<not>bval b (s \<circ> venv e);  e \<turnstile> (c\<^isub>2,s) \<Rightarrow> t \<rbrakk> \<Longrightarrow>
       
    26          e \<turnstile> (IF b THEN c\<^isub>1 ELSE c\<^isub>2, s) \<Rightarrow> t" |
       
    27 
       
    28 WhileFalse: "\<not>bval b (s \<circ> venv e) \<Longrightarrow> e \<turnstile> (WHILE b DO c,s) \<Rightarrow> s" |
       
    29 WhileTrue:
       
    30   "\<lbrakk> bval b (s\<^isub>1 \<circ> venv e);  e \<turnstile> (c,s\<^isub>1) \<Rightarrow> s\<^isub>2;
       
    31      e \<turnstile> (WHILE b DO c, s\<^isub>2) \<Rightarrow> s\<^isub>3 \<rbrakk> \<Longrightarrow>
       
    32    e \<turnstile> (WHILE b DO c, s\<^isub>1) \<Rightarrow> s\<^isub>3" |
       
    33 
       
    34 Var: "(pe,ve(x:=f),f+1) \<turnstile> (c,s) \<Rightarrow> t  \<Longrightarrow>
       
    35       (pe,ve,f) \<turnstile> ({VAR x;; c}, s) \<Rightarrow> t(f := s f)" |
       
    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
       
    43       \<Longrightarrow>  (pe,ve,f) \<turnstile> ({PROC p = cp;; c}, s) \<Rightarrow> t"
       
    44 
       
    45 code_pred big_step .
       
    46 
       
    47 values "{map t [0,1] |t. ([], \<lambda>n. 0, 0) \<turnstile> (CALL ''p'', nth [42, 43]) \<Rightarrow> t}"
       
    48 
       
    49 (* FIXME: syntax *)
       
    50 values "{map t [0, 1, 2] |t. ([], (\<lambda>_. 3)(''x'' := 0, ''y'' := 1,''z'' := 2), 0) \<turnstile> (test_com, (%_. 0)) \<Rightarrow> t}"
       
    51 
       
    52 
       
    53 end