| 11229 |      1 | (*  Title:      HOL/MicroJava/BV/Typing_Framework.thy
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Tobias Nipkow
 | 
|  |      4 |     Copyright   2000 TUM
 | 
|  |      5 | *)
 | 
|  |      6 | 
 | 
| 12911 |      7 | header {* \isaheader{Typing and Dataflow Analysis Framework} *}
 | 
| 11229 |      8 | 
 | 
| 16417 |      9 | theory Typing_Framework imports Listn begin
 | 
| 11229 |     10 | 
 | 
|  |     11 | text {* 
 | 
| 13224 |     12 |   The relationship between dataflow analysis and a welltyped-instruction predicate. 
 | 
| 11229 |     13 | *}
 | 
| 12516 |     14 | types
 | 
|  |     15 |   's step_type = "nat \<Rightarrow> 's \<Rightarrow> (nat \<times> 's) list"
 | 
| 11229 |     16 | 
 | 
|  |     17 | constdefs
 | 
| 13006 |     18 |  stable :: "'s ord \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> nat \<Rightarrow> bool"
 | 
| 12516 |     19 | "stable r step ss p == !(q,s'):set(step p (ss!p)). s' <=_r ss!q"
 | 
| 11229 |     20 | 
 | 
| 13006 |     21 |  stables :: "'s ord \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> bool"
 | 
| 12516 |     22 | "stables r step ss == !p<size ss. stable r step ss p"
 | 
| 11229 |     23 | 
 | 
| 14653 |     24 |  wt_step ::
 | 
|  |     25 | "'s ord \<Rightarrow> 's \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> bool"
 | 
|  |     26 | "wt_step r T step ts ==
 | 
|  |     27 |  !p<size(ts). ts!p ~= T & stable r step ts p"
 | 
|  |     28 | 
 | 
| 13006 |     29 |  is_bcv :: "'s ord \<Rightarrow> 's \<Rightarrow> 's step_type 
 | 
|  |     30 |            \<Rightarrow> nat \<Rightarrow> 's set \<Rightarrow> ('s list \<Rightarrow> 's list) \<Rightarrow> bool"  
 | 
| 12516 |     31 | "is_bcv r T step n A bcv == !ss : list n A.
 | 
| 11229 |     32 |    (!p<n. (bcv ss)!p ~= T) =
 | 
| 12516 |     33 |    (? ts: list n A. ss <=[r] ts & wt_step r T step ts)"
 | 
| 11229 |     34 | 
 | 
|  |     35 | end
 |