src/HOL/IMP/Abs_Int_ITP/Abs_State_ITP.thy
changeset 64448 49b78f1f9e01
parent 64447 e44f5c123f26
child 64449 8c44dfb4ca8a
equal deleted inserted replaced
64447:e44f5c123f26 64448:49b78f1f9e01
     1 (* Author: Tobias Nipkow *)
       
     2 
       
     3 theory Abs_State_ITP
       
     4 imports Abs_Int0_ITP
       
     5   "~~/src/HOL/Library/Char_ord" "~~/src/HOL/Library/List_lexord"
       
     6   (* Library import merely to allow string lists to be sorted for output *)
       
     7 begin
       
     8 
       
     9 subsection "Abstract State with Computable Ordering"
       
    10 
       
    11 text{* A concrete type of state with computable @{text"\<sqsubseteq>"}: *}
       
    12 
       
    13 datatype 'a st = FunDom "vname \<Rightarrow> 'a" "vname list"
       
    14 
       
    15 fun "fun" where "fun (FunDom f xs) = f"
       
    16 fun dom where "dom (FunDom f xs) = xs"
       
    17 
       
    18 definition [simp]: "inter_list xs ys = [x\<leftarrow>xs. x \<in> set ys]"
       
    19 
       
    20 definition "show_st S = [(x,fun S x). x \<leftarrow> sort(dom S)]"
       
    21 
       
    22 definition "show_acom = map_acom (map_option show_st)"
       
    23 definition "show_acom_opt = map_option show_acom"
       
    24 
       
    25 definition "lookup F x = (if x : set(dom F) then fun F x else \<top>)"
       
    26 
       
    27 definition "update F x y =
       
    28   FunDom ((fun F)(x:=y)) (if x \<in> set(dom F) then dom F else x # dom F)"
       
    29 
       
    30 lemma lookup_update: "lookup (update S x y) = (lookup S)(x:=y)"
       
    31 by(rule ext)(auto simp: lookup_def update_def)
       
    32 
       
    33 definition "\<gamma>_st \<gamma> F = {f. \<forall>x. f x \<in> \<gamma>(lookup F x)}"
       
    34 
       
    35 instantiation st :: (SL_top) SL_top
       
    36 begin
       
    37 
       
    38 definition "le_st F G = (ALL x : set(dom G). lookup F x \<sqsubseteq> fun G x)"
       
    39 
       
    40 definition
       
    41 "join_st F G =
       
    42  FunDom (\<lambda>x. fun F x \<squnion> fun G x) (inter_list (dom F) (dom G))"
       
    43 
       
    44 definition "\<top> = FunDom (\<lambda>x. \<top>) []"
       
    45 
       
    46 instance
       
    47 proof
       
    48   case goal2 thus ?case
       
    49     apply(auto simp: le_st_def)
       
    50     by (metis lookup_def preord_class.le_trans top)
       
    51 qed (auto simp: le_st_def lookup_def join_st_def Top_st_def)
       
    52 
       
    53 end
       
    54 
       
    55 lemma mono_lookup: "F \<sqsubseteq> F' \<Longrightarrow> lookup F x \<sqsubseteq> lookup F' x"
       
    56 by(auto simp add: lookup_def le_st_def)
       
    57 
       
    58 lemma mono_update: "a \<sqsubseteq> a' \<Longrightarrow> S \<sqsubseteq> S' \<Longrightarrow> update S x a \<sqsubseteq> update S' x a'"
       
    59 by(auto simp add: le_st_def lookup_def update_def)
       
    60 
       
    61 locale Gamma = Val_abs where \<gamma>=\<gamma> for \<gamma> :: "'av::SL_top \<Rightarrow> val set"
       
    62 begin
       
    63 
       
    64 abbreviation \<gamma>\<^sub>f :: "'av st \<Rightarrow> state set"
       
    65 where "\<gamma>\<^sub>f == \<gamma>_st \<gamma>"
       
    66 
       
    67 abbreviation \<gamma>\<^sub>o :: "'av st option \<Rightarrow> state set"
       
    68 where "\<gamma>\<^sub>o == \<gamma>_option \<gamma>\<^sub>f"
       
    69 
       
    70 abbreviation \<gamma>\<^sub>c :: "'av st option acom \<Rightarrow> state set acom"
       
    71 where "\<gamma>\<^sub>c == map_acom \<gamma>\<^sub>o"
       
    72 
       
    73 lemma gamma_f_Top[simp]: "\<gamma>\<^sub>f Top = UNIV"
       
    74 by(auto simp: Top_st_def \<gamma>_st_def lookup_def)
       
    75 
       
    76 lemma gamma_o_Top[simp]: "\<gamma>\<^sub>o Top = UNIV"
       
    77 by (simp add: Top_option_def)
       
    78 
       
    79 (* FIXME (maybe also le \<rightarrow> sqle?) *)
       
    80 
       
    81 lemma mono_gamma_f: "f \<sqsubseteq> g \<Longrightarrow> \<gamma>\<^sub>f f \<subseteq> \<gamma>\<^sub>f g"
       
    82 apply(simp add:\<gamma>_st_def subset_iff lookup_def le_st_def split: if_splits)
       
    83 by (metis UNIV_I mono_gamma gamma_Top subsetD)
       
    84 
       
    85 lemma mono_gamma_o:
       
    86   "sa \<sqsubseteq> sa' \<Longrightarrow> \<gamma>\<^sub>o sa \<subseteq> \<gamma>\<^sub>o sa'"
       
    87 by(induction sa sa' rule: le_option.induct)(simp_all add: mono_gamma_f)
       
    88 
       
    89 lemma mono_gamma_c: "ca \<sqsubseteq> ca' \<Longrightarrow> \<gamma>\<^sub>c ca \<le> \<gamma>\<^sub>c ca'"
       
    90 by (induction ca ca' rule: le_acom.induct) (simp_all add:mono_gamma_o)
       
    91 
       
    92 lemma in_gamma_option_iff:
       
    93   "x : \<gamma>_option r u \<longleftrightarrow> (\<exists>u'. u = Some u' \<and> x : r u')"
       
    94 by (cases u) auto
       
    95 
       
    96 end
       
    97 
       
    98 end