| 43158 |      1 | (* Author: Tobias Nipkow *)
 | 
|  |      2 | 
 | 
|  |      3 | theory Def_Ass_Big imports Com Def_Ass_Exp
 | 
|  |      4 | begin
 | 
|  |      5 | 
 | 
|  |      6 | subsection "Initialization-Sensitive Big Step Semantics"
 | 
|  |      7 | 
 | 
|  |      8 | inductive
 | 
|  |      9 |   big_step :: "(com \<times> state option) \<Rightarrow> state option \<Rightarrow> bool" (infix "\<Rightarrow>" 55)
 | 
|  |     10 | where
 | 
|  |     11 | None: "(c,None) \<Rightarrow> None" |
 | 
|  |     12 | Skip: "(SKIP,s) \<Rightarrow> s" |
 | 
|  |     13 | AssignNone: "aval a s = None \<Longrightarrow> (x ::= a, Some s) \<Rightarrow> None" |
 | 
|  |     14 | Assign: "aval a s = Some i \<Longrightarrow> (x ::= a, Some s) \<Rightarrow> Some(s(x := Some i))" |
 | 
| 47818 |     15 | Seq:    "(c\<^isub>1,s\<^isub>1) \<Rightarrow> s\<^isub>2 \<Longrightarrow> (c\<^isub>2,s\<^isub>2) \<Rightarrow> s\<^isub>3 \<Longrightarrow> (c\<^isub>1;c\<^isub>2,s\<^isub>1) \<Rightarrow> s\<^isub>3" |
 | 
| 43158 |     16 | 
 | 
|  |     17 | IfNone:  "bval b s = None \<Longrightarrow> (IF b THEN c\<^isub>1 ELSE c\<^isub>2,Some s) \<Rightarrow> None" |
 | 
|  |     18 | IfTrue:  "\<lbrakk> bval b s = Some True;  (c\<^isub>1,Some s) \<Rightarrow> s' \<rbrakk> \<Longrightarrow>
 | 
|  |     19 |   (IF b THEN c\<^isub>1 ELSE c\<^isub>2,Some s) \<Rightarrow> s'" |
 | 
|  |     20 | IfFalse: "\<lbrakk> bval b s = Some False;  (c\<^isub>2,Some s) \<Rightarrow> s' \<rbrakk> \<Longrightarrow>
 | 
|  |     21 |   (IF b THEN c\<^isub>1 ELSE c\<^isub>2,Some s) \<Rightarrow> s'" |
 | 
|  |     22 | 
 | 
|  |     23 | WhileNone: "bval b s = None \<Longrightarrow> (WHILE b DO c,Some s) \<Rightarrow> None" |
 | 
|  |     24 | WhileFalse: "bval b s = Some False \<Longrightarrow> (WHILE b DO c,Some s) \<Rightarrow> Some s" |
 | 
|  |     25 | WhileTrue:
 | 
|  |     26 |   "\<lbrakk> bval b s = Some True;  (c,Some s) \<Rightarrow> s';  (WHILE b DO c,s') \<Rightarrow> s'' \<rbrakk> \<Longrightarrow>
 | 
|  |     27 |   (WHILE b DO c,Some s) \<Rightarrow> s''"
 | 
|  |     28 | 
 | 
|  |     29 | lemmas big_step_induct = big_step.induct[split_format(complete)]
 | 
|  |     30 | 
 | 
|  |     31 | end
 |