# HG changeset patch # User nipkow # Date 1334911615 -7200 # Node ID 0d3e95375bb72408e5bade135f99041b654ce265 # Parent 540a5af9a01c41c899d8493cfdf8977b283ebd35 forgot to add file diff -r 540a5af9a01c -r 0d3e95375bb7 src/HOL/IMP/Abs_State.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/IMP/Abs_State.thy Fri Apr 20 10:46:55 2012 +0200 @@ -0,0 +1,213 @@ +(* Author: Tobias Nipkow *) + +theory Abs_State +imports Abs_Int0 +begin + +subsubsection "Welltypedness: wt" + +instantiation com :: vars +begin + +fun vars_com :: "com \ vname set" where +"vars com.SKIP = {}" | +"vars (x::=e) = {x} \ vars e" | +"vars (c1;c2) = vars c1 \ vars c2" | +"vars (IF b THEN c1 ELSE c2) = vars b \ vars c1 \ vars c2" | +"vars (WHILE b DO c) = vars b \ vars c" + +instance .. + +end + + +lemma finite_avars: "finite(vars(a::aexp))" +by(induction a) simp_all + +lemma finite_bvars: "finite(vars(b::bexp))" +by(induction b) (simp_all add: finite_avars) + +lemma finite_cvars: "finite(vars(c::com))" +by(induction c) (simp_all add: finite_avars finite_bvars) + + +class wt = +fixes wt :: "'a \ vname set \ bool" + + +instantiation acom :: (wt)wt +begin + +definition wt_acom where +"wt C X = (vars(strip C) \ X \ (\a \ set(annos C). wt a X))" + +instance .. + +end + + +instantiation option :: (wt)wt +begin + +definition wt_option where +"wt opt X = (case opt of None \ True | Some x \ wt x X)" + +lemma wt_simps[simp]: "wt None X" "wt (Some x) X = wt x X" +by(simp_all add: wt_option_def) + +instance .. + +end + +class SL_top_wt = join + wt + +fixes top :: "com \ 'a" ("\\<^bsub>_\<^esub>") +assumes join_ge1 [simp]: "wt x X \ wt y X \ x \ x \ y" +and join_ge2 [simp]: "wt x X \ wt y X \ y \ x \ y" +and join_least[simp]: "x \ z \ y \ z \ x \ y \ z" +and top[simp]: "wt x (vars c) \ x \ top c" +and wt_top[simp]: "wt (top c) (vars c)" +and wt_join[simp]: "wt x X \ wt y X \ wt (x \ y) X" + + +instantiation option :: (SL_top_wt)SL_top_wt +begin + +definition top_option where "top c = Some(top c)" + +instance proof + case goal1 thus ?case by(cases x, simp, cases y, simp_all) +next + case goal2 thus ?case by(cases y, simp, cases x, simp_all) +next + case goal3 thus ?case by(cases z, simp, cases y, simp, cases x, simp_all) +next + case goal4 thus ?case by(cases x, simp_all add: top_option_def) +next + case goal5 thus ?case by(simp add: top_option_def) +next + case goal6 thus ?case by(simp add: wt_option_def split: option.splits) +qed + +end + + +subsection "Abstract State with Computable Ordering" + +hide_type st --"to avoid long names" + +text{* A concrete type of state with computable @{text"\"}: *} + +datatype 'a st = FunDom "vname \ 'a" "vname set" + +fun "fun" where "fun (FunDom f xs) = f" +fun dom where "dom (FunDom f xs) = xs" + +definition "show_st S = (\x. (x, fun S x)) ` dom S" + +value [code] "show_st (FunDom (\x. 1::int) {''a'',''b''})" + +definition "show_acom = map_acom (Option.map show_st)" +definition "show_acom_opt = Option.map show_acom" + +definition "update F x y = FunDom ((fun F)(x:=y)) (dom F)" + +lemma fun_update[simp]: "fun (update S x y) = (fun S)(x:=y)" +by(rule ext)(auto simp: update_def) + +lemma dom_update[simp]: "dom (update S x y) = dom S" +by(simp add: update_def) + +definition "\_st \ F = {f. \x\dom F. f x \ \(fun F x)}" + + +instantiation st :: (preord) preord +begin + +definition le_st :: "'a st \ 'a st \ bool" where +"F \ G = (dom F = dom G \ (\x \ dom F. fun F x \ fun G x))" + +instance +proof + case goal2 thus ?case by(auto simp: le_st_def)(metis preord_class.le_trans) +qed (auto simp: le_st_def) + +end + +instantiation st :: (join) join +begin + +definition join_st :: "'a st \ 'a st \ 'a st" where +"F \ G = FunDom (\x. fun F x \ fun G x) (dom F)" + +instance .. + +end + +instantiation st :: (type) wt +begin + +definition wt_st :: "'a st \ vname set \ bool" where +"wt F X = (dom F = X)" + +instance .. + +end + +instantiation st :: (SL_top) SL_top_wt +begin + +definition top_st where "top c = FunDom (\x. \) (vars c)" + +instance +proof +qed (auto simp: le_st_def join_st_def top_st_def wt_st_def) + +end + + +lemma mono_fun: "F \ G \ x : dom F \ fun F x \ fun G x" +by(auto simp: le_st_def) + +lemma mono_update[simp]: + "a1 \ a2 \ S1 \ S2 \ update S1 x a1 \ update S2 x a2" +by(auto simp add: le_st_def update_def) + + +locale Gamma = Val_abs where \=\ for \ :: "'av::SL_top \ val set" +begin + +abbreviation \\<^isub>f :: "'av st \ state set" +where "\\<^isub>f == \_st \" + +abbreviation \\<^isub>o :: "'av st option \ state set" +where "\\<^isub>o == \_option \\<^isub>f" + +abbreviation \\<^isub>c :: "'av st option acom \ state set acom" +where "\\<^isub>c == map_acom \\<^isub>o" + +lemma gamma_f_Top[simp]: "\\<^isub>f (top c) = UNIV" +by(auto simp: top_st_def \_st_def) + +lemma gamma_o_Top[simp]: "\\<^isub>o (top c) = UNIV" +by (simp add: top_option_def) + +(* FIXME (maybe also le \ sqle?) *) + +lemma mono_gamma_f: "f \ g \ \\<^isub>f f \ \\<^isub>f g" +apply(simp add:\_st_def subset_iff le_st_def split: if_splits) +by (metis mono_gamma subsetD) + +lemma mono_gamma_o: + "S1 \ S2 \ \\<^isub>o S1 \ \\<^isub>o S2" +by(induction S1 S2 rule: le_option.induct)(simp_all add: mono_gamma_f) + +lemma mono_gamma_c: "C1 \ C2 \ \\<^isub>c C1 \ \\<^isub>c C2" +by (induction C1 C2 rule: le_acom.induct) (simp_all add:mono_gamma_o) + +lemma in_gamma_option_iff: + "x : \_option r u \ (\u'. u = Some u' \ x : r u')" +by (cases u) auto + +end + +end