| author | wenzelm | 
| Mon, 01 May 2017 11:04:33 +0200 | |
| changeset 65659 | 293141fb093d | 
| parent 53015 | a1119cf551e8 | 
| child 67406 | 23307fd33906 | 
| permissions | -rw-r--r-- | 
| 43158 | 1  | 
(* Author: Tobias Nipkow *)  | 
2  | 
||
| 50161 | 3  | 
theory Def_Init_Big  | 
| 52726 | 4  | 
imports Def_Init_Exp Def_Init  | 
| 43158 | 5  | 
begin  | 
6  | 
||
7  | 
subsection "Initialization-Sensitive Big Step Semantics"  | 
|
8  | 
||
9  | 
inductive  | 
|
10  | 
big_step :: "(com \<times> state option) \<Rightarrow> state option \<Rightarrow> bool" (infix "\<Rightarrow>" 55)  | 
|
11  | 
where  | 
|
12  | 
None: "(c,None) \<Rightarrow> None" |  | 
|
13  | 
Skip: "(SKIP,s) \<Rightarrow> s" |  | 
|
14  | 
AssignNone: "aval a s = None \<Longrightarrow> (x ::= a, Some s) \<Rightarrow> None" |  | 
|
15  | 
Assign: "aval a s = Some i \<Longrightarrow> (x ::= a, Some s) \<Rightarrow> Some(s(x := Some i))" |  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52726 
diff
changeset
 | 
16  | 
Seq: "(c\<^sub>1,s\<^sub>1) \<Rightarrow> s\<^sub>2 \<Longrightarrow> (c\<^sub>2,s\<^sub>2) \<Rightarrow> s\<^sub>3 \<Longrightarrow> (c\<^sub>1;;c\<^sub>2,s\<^sub>1) \<Rightarrow> s\<^sub>3" |  | 
| 43158 | 17  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52726 
diff
changeset
 | 
18  | 
IfNone: "bval b s = None \<Longrightarrow> (IF b THEN c\<^sub>1 ELSE c\<^sub>2,Some s) \<Rightarrow> None" |  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52726 
diff
changeset
 | 
19  | 
IfTrue: "\<lbrakk> bval b s = Some True; (c\<^sub>1,Some s) \<Rightarrow> s' \<rbrakk> \<Longrightarrow>  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52726 
diff
changeset
 | 
20  | 
(IF b THEN c\<^sub>1 ELSE c\<^sub>2,Some s) \<Rightarrow> s'" |  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52726 
diff
changeset
 | 
21  | 
IfFalse: "\<lbrakk> bval b s = Some False; (c\<^sub>2,Some s) \<Rightarrow> s' \<rbrakk> \<Longrightarrow>  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52726 
diff
changeset
 | 
22  | 
(IF b THEN c\<^sub>1 ELSE c\<^sub>2,Some s) \<Rightarrow> s'" |  | 
| 43158 | 23  | 
|
24  | 
WhileNone: "bval b s = None \<Longrightarrow> (WHILE b DO c,Some s) \<Rightarrow> None" |  | 
|
25  | 
WhileFalse: "bval b s = Some False \<Longrightarrow> (WHILE b DO c,Some s) \<Rightarrow> Some s" |  | 
|
26  | 
WhileTrue:  | 
|
27  | 
"\<lbrakk> bval b s = Some True; (c,Some s) \<Rightarrow> s'; (WHILE b DO c,s') \<Rightarrow> s'' \<rbrakk> \<Longrightarrow>  | 
|
28  | 
(WHILE b DO c,Some s) \<Rightarrow> s''"  | 
|
29  | 
||
30  | 
lemmas big_step_induct = big_step.induct[split_format(complete)]  | 
|
31  | 
||
| 52726 | 32  | 
|
33  | 
subsection "Soundness wrt Big Steps"  | 
|
34  | 
||
35  | 
text{* Note the special form of the induction because one of the arguments
 | 
|
36  | 
of the inductive predicate is not a variable but the term @{term"Some s"}: *}
 | 
|
37  | 
||
38  | 
theorem Sound:  | 
|
39  | 
"\<lbrakk> (c,Some s) \<Rightarrow> s'; D A c A'; A \<subseteq> dom s \<rbrakk>  | 
|
40  | 
\<Longrightarrow> \<exists> t. s' = Some t \<and> A' \<subseteq> dom t"  | 
|
41  | 
proof (induction c "Some s" s' arbitrary: s A A' rule:big_step_induct)  | 
|
42  | 
case AssignNone thus ?case  | 
|
43  | 
by auto (metis aval_Some option.simps(3) subset_trans)  | 
|
44  | 
next  | 
|
45  | 
case Seq thus ?case by auto metis  | 
|
46  | 
next  | 
|
47  | 
case IfTrue thus ?case by auto blast  | 
|
48  | 
next  | 
|
49  | 
case IfFalse thus ?case by auto blast  | 
|
50  | 
next  | 
|
51  | 
case IfNone thus ?case  | 
|
52  | 
by auto (metis bval_Some option.simps(3) order_trans)  | 
|
53  | 
next  | 
|
54  | 
case WhileNone thus ?case  | 
|
55  | 
by auto (metis bval_Some option.simps(3) order_trans)  | 
|
56  | 
next  | 
|
57  | 
case (WhileTrue b s c s' s'')  | 
|
58  | 
from `D A (WHILE b DO c) A'` obtain A' where "D A c A'" by blast  | 
|
59  | 
then obtain t' where "s' = Some t'" "A \<subseteq> dom t'"  | 
|
60  | 
by (metis D_incr WhileTrue(3,7) subset_trans)  | 
|
61  | 
from WhileTrue(5)[OF this(1) WhileTrue(6) this(2)] show ?case .  | 
|
62  | 
qed auto  | 
|
63  | 
||
64  | 
corollary sound: "\<lbrakk> D (dom s) c A'; (c,Some s) \<Rightarrow> s' \<rbrakk> \<Longrightarrow> s' \<noteq> None"  | 
|
65  | 
by (metis Sound not_Some_eq subset_refl)  | 
|
66  | 
||
| 43158 | 67  | 
end  | 
| 52726 | 68  |