src/HOL/MicroJava/BV/Typing_Framework.thy
author wenzelm
Wed, 13 Jul 2005 16:07:23 +0200
changeset 16802 6eeee59dac4c
parent 16417 9bc16273c2d4
child 27680 5a557a5afc48
permissions -rw-r--r--
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook; added call_atp: bool ref; do 'setmp print_mode []', which is more robust than manual ref manipulation; added subtract_simpset, subtract_claset (supercede delta approximation);

(*  Title:      HOL/MicroJava/BV/Typing_Framework.thy
    ID:         $Id$
    Author:     Tobias Nipkow
    Copyright   2000 TUM
*)

header {* \isaheader{Typing and Dataflow Analysis Framework} *}

theory Typing_Framework imports Listn begin

text {* 
  The relationship between dataflow analysis and a welltyped-instruction predicate. 
*}
types
  's step_type = "nat \<Rightarrow> 's \<Rightarrow> (nat \<times> 's) list"

constdefs
 stable :: "'s ord \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> nat \<Rightarrow> bool"
"stable r step ss p == !(q,s'):set(step p (ss!p)). s' <=_r ss!q"

 stables :: "'s ord \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> bool"
"stables r step ss == !p<size ss. stable r step ss p"

 wt_step ::
"'s ord \<Rightarrow> 's \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> bool"
"wt_step r T step ts ==
 !p<size(ts). ts!p ~= T & stable r step ts p"

 is_bcv :: "'s ord \<Rightarrow> 's \<Rightarrow> 's step_type 
           \<Rightarrow> nat \<Rightarrow> 's set \<Rightarrow> ('s list \<Rightarrow> 's list) \<Rightarrow> bool"  
"is_bcv r T step n A bcv == !ss : list n A.
   (!p<n. (bcv ss)!p ~= T) =
   (? ts: list n A. ss <=[r] ts & wt_step r T step ts)"

end