# HG changeset patch # User nipkow # Date 1314984344 -7200 # Node ID 17dbd9d9db380a47fe94a108036fac7ce1216a3e # Parent fe03653315667519fd9f54b81451a6e81a307e53# Parent 22bbd0d1b943422ac0653dba14a0aef5b7dc9c21 merged diff -r fe0365331566 -r 17dbd9d9db38 src/HOL/IMP/AbsInt0.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/IMP/AbsInt0.thy Fri Sep 02 19:25:44 2011 +0200 @@ -0,0 +1,66 @@ +(* Author: Tobias Nipkow *) + +theory AbsInt0 +imports Astate +begin + +subsection "Computable Abstract Interpretation" + +text{* Abstract interpretation over type @{text astate} instead of +functions. *} + +locale Abs_Int = Val_abs +begin + +fun aval' :: "aexp \ 'a astate \ 'a" ("aval\<^isup>#") where +"aval' (N n) _ = num' n" | +"aval' (V x) S = lookup S x" | +"aval' (Plus e1 e2) S = plus' (aval' e1 S) (aval' e2 S)" + +abbreviation astate_in_rep (infix "<:" 50) where +"s <: S == ALL x. s x <: lookup S x" + +lemma astate_in_rep_le: "(s::state) <: S \ S \ T \ s <: T" +by (metis in_mono le_astate_def le_rep lookup_def top) + +lemma aval'_sound: "s <: S \ aval a s <: aval' a S" +by (induct a) (auto simp: rep_num' rep_plus') + + +fun AI :: "com \ 'a astate \ 'a astate" where +"AI SKIP S = S" | +"AI (x ::= a) S = update S x (aval' a S)" | +"AI (c1;c2) S = AI c2 (AI c1 S)" | +"AI (IF b THEN c1 ELSE c2) S = (AI c1 S) \ (AI c2 S)" | +"AI (WHILE b DO c) S = pfp_above (AI c) S" + +lemma AI_sound: "(c,s) \ t \ s <: S0 \ t <: AI c S0" +proof(induct c arbitrary: s t S0) + case SKIP thus ?case by fastsimp +next + case Assign thus ?case + by (auto simp: lookup_update aval'_sound) +next + case Semi thus ?case by auto +next + case If thus ?case + by (metis AI.simps(4) IfE astate_in_rep_le join_ge1 join_ge2) +next + case (While b c) + let ?P = "pfp_above (AI c) S0" + have pfp: "AI c ?P \ ?P" and "S0 \ ?P" + by(simp_all add: SL_top_class.pfp_above_pfp) + { fix s t have "(WHILE b DO c,s) \ t \ s <: ?P \ t <: ?P" + proof(induct "WHILE b DO c" s t rule: big_step_induct) + case WhileFalse thus ?case by simp + next + case WhileTrue thus ?case using While.hyps pfp astate_in_rep_le by metis + qed + } + with astate_in_rep_le[OF `s <: S0` `S0 \ ?P`] + show ?case by (metis While(2) AI.simps(5)) +qed + +end + +end diff -r fe0365331566 -r 17dbd9d9db38 src/HOL/IMP/AbsInt0_const.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/IMP/AbsInt0_const.thy Fri Sep 02 19:25:44 2011 +0200 @@ -0,0 +1,110 @@ +(* Author: Tobias Nipkow *) + +theory AbsInt0_const +imports AbsInt0 +begin + +subsection "Constant Propagation" + +datatype cval = Const val | Any + +fun rep_cval where +"rep_cval (Const n) = {n}" | +"rep_cval (Any) = UNIV" + +fun plus_cval where +"plus_cval (Const m) (Const n) = Const(m+n)" | +"plus_cval _ _ = Any" + +instantiation cval :: SL_top +begin + +fun le_cval where +"_ \ Any = True" | +"Const n \ Const m = (n=m)" | +"Any \ Const _ = False" + +fun join_cval where +"Const m \ Const n = (if n=m then Const m else Any)" | +"_ \ _ = Any" + +definition "Top = Any" + +instance +proof + case goal1 thus ?case by (cases x) simp_all +next + case goal2 thus ?case by(cases z, cases y, cases x, simp_all) +next + case goal3 thus ?case by(cases x, cases y, simp_all) +next + case goal4 thus ?case by(cases y, cases x, simp_all) +next + case goal5 thus ?case by(cases z, cases y, cases x, simp_all) +next + case goal6 thus ?case by(simp add: Top_cval_def) +qed + +end + +interpretation Rep rep_cval +proof + case goal1 thus ?case + by(cases a, cases b, simp, simp, cases b, simp, simp) +qed + +interpretation Val_abs rep_cval Const plus_cval +proof + case goal1 show ?case by simp +next + case goal2 thus ?case + by(cases a1, cases a2, simp, simp, cases a2, simp, simp) +qed + +interpretation Abs_Int rep_cval Const plus_cval +defines AI_const is AI +and aval'_const is aval' +.. + +text{* Straight line code: *} +definition "test1_const = + ''y'' ::= N 7; + ''z'' ::= Plus (V ''y'') (N 2); + ''y'' ::= Plus (V ''x'') (N 0)" +value [code] "list (AI_const test1_const Top)" + +text{* Conditional: *} +definition "test2_const = + IF Less (N 41) (V ''x'') THEN ''x'' ::= N 5 ELSE ''x'' ::= N 5" +value "list (AI_const test2_const Top)" + +text{* Conditional, test is ignored: *} +definition "test3_const = + ''x'' ::= N 42; + IF Less (N 41) (V ''x'') THEN ''x'' ::= N 5 ELSE ''x'' ::= N 6" +value "list (AI_const test3_const Top)" + +text{* While: *} +definition "test4_const = + ''x'' ::= N 0; WHILE B True DO ''x'' ::= N 0" +value [code] "list (AI_const test4_const Top)" + +text{* While, test is ignored: *} +definition "test5_const = + ''x'' ::= N 0; WHILE Less (V ''x'') (N 1) DO ''x'' ::= N 1" +value [code] "list (AI_const test5_const Top)" + +text{* Iteration is needed: *} +definition "test6_const = + ''x'' ::= N 0; ''y'' ::= N 0; ''z'' ::= N 2; + WHILE Less (V ''x'') (N 1) DO (''x'' ::= V ''y''; ''y'' ::= V ''z'')" +value [code] "list (AI_const test6_const Top)" + +text{* More iteration would be needed: *} +definition "test7_const = + ''x'' ::= N 0; ''y'' ::= N 0; ''z'' ::= N 0; ''u'' ::= N 3; + WHILE B True DO (''x'' ::= V ''y''; ''y'' ::= V ''z''; ''z'' ::= V ''u'')" +value [code] "list (AI_const test7_const Top)" + + +end diff -r fe0365331566 -r 17dbd9d9db38 src/HOL/IMP/AbsInt0_fun.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/IMP/AbsInt0_fun.thy Fri Sep 02 19:25:44 2011 +0200 @@ -0,0 +1,213 @@ +(* Author: Tobias Nipkow *) + +header "Abstract Interpretation" + +theory AbsInt0_fun +imports "~~/src/HOL/ex/Interpretation_with_Defs" Big_Step +begin + +subsection "Orderings" + +class preord = +fixes le :: "'a \ 'a \ bool" (infix "\" 50) +assumes le_refl[simp]: "x \ x" +and le_trans: "x \ y \ y \ z \ x \ z" + +text{* Note: no antisymmetry. Allows implementations where some abstract +element is implemented by two different values @{prop "x \ y"} +such that @{prop"x \ y"} and @{prop"y \ x"}. Antisymmetry is not +needed because we never compare elements for equality but only for @{text"\"}. +*} + +class SL_top = preord + +fixes join :: "'a \ 'a \ 'a" (infixl "\" 65) +fixes Top :: "'a" +assumes join_ge1 [simp]: "x \ x \ y" +and join_ge2 [simp]: "y \ x \ y" +and join_least: "x \ z \ y \ z \ x \ y \ z" +and top[simp]: "x \ Top" +begin + +lemma join_le_iff[simp]: "x \ y \ z \ x \ z \ y \ z" +by (metis join_ge1 join_ge2 join_least le_trans) + +fun bpfp where +"bpfp f 0 _ = Top" | +"bpfp f (Suc n) x = (if f x \ x then x else bpfp f n (f x))" + +definition "pfp f = bpfp f 3" + +lemma postfixedpoint: "f(bpfp f n x) \ bpfp f n x" +apply (induct n arbitrary: x) + apply (simp) +apply (simp) +done + +lemma bpfp_funpow: "bpfp f n x \ Top \ EX k. bpfp f n x = (f^^k) x" +apply(induct n arbitrary: x) +apply simp +apply simp +apply (auto) +apply(rule_tac x=0 in exI) +apply simp +by (metis funpow.simps(2) funpow_swap1 o_apply) + +lemma pfp_funpow: "pfp f x \ Top \ EX k. pfp f x = (f^^k) x" +by(simp add: pfp_def bpfp_funpow) + +abbreviation (input) lift (infix "\" 65) where "f\x0 == (%x. x0 \ f x)" + +definition "pfp_above f x0 = pfp (f\x0) x0" + +lemma pfp_above_pfp: +shows "f(pfp_above f x0) \ pfp_above f x0" and "x0 \ pfp_above f x0" +using postfixedpoint[of "lift f x0"] +by(auto simp add: pfp_above_def pfp_def) + +lemma least_pfp: +assumes mono: "!!x y. x \ y \ f x \ f y" and "pfp_above f x0 \ Top" +and "f p \ p" and "x0 \ p" shows "pfp_above f x0 \ p" +proof- + let ?F = "lift f x0" + obtain k where "pfp_above f x0 = (?F^^k) x0" + using pfp_funpow `pfp_above f x0 \ Top` + by(fastsimp simp add: pfp_above_def) + moreover + { fix n have "(?F^^n) x0 \ p" + proof(induct n) + case 0 show ?case by(simp add: `x0 \ p`) + next + case (Suc n) thus ?case + by (simp add: `x0 \ p`)(metis Suc assms(3) le_trans mono) + qed + } ultimately show ?thesis by simp +qed + +lemma chain: assumes "!!x y. x \ y \ f x \ f y" +shows "((f\x0)^^i) x0 \ ((f\x0)^^(i+1)) x0" +apply(induct i) + apply simp +apply simp +by (metis assms join_ge2 le_trans) + +lemma pfp_almost_fp: +assumes mono: "!!x y. x \ y \ f x \ f y" and "pfp_above f x0 \ Top" +shows "pfp_above f x0 \ x0 \ f(pfp_above f x0)" +proof- + let ?p = "pfp_above f x0" + obtain k where 1: "?p = ((f\x0)^^k) x0" + using pfp_funpow `?p \ Top` + by(fastsimp simp add: pfp_above_def) + thus ?thesis using chain mono by simp +qed + +end + +text{* The interface of abstract values: *} + +locale Rep = +fixes rep :: "'a::SL_top \ 'b set" +assumes le_rep: "a \ b \ rep a \ rep b" +begin + +abbreviation in_rep (infix "<:" 50) where "x <: a == x : rep a" + +lemma in_rep_join: "x <: a1 \ x <: a2 \ x <: a1 \ a2" +by (metis SL_top_class.join_ge1 SL_top_class.join_ge2 le_rep subsetD) + +end + +locale Val_abs = Rep rep + for rep :: "'a::SL_top \ val set" + +fixes num' :: "val \ 'a" ("num\<^isup>#") +and plus' :: "'a \ 'a \ 'a" ("plus\<^isup>#") +assumes rep_num': "rep(num' n) = {n}" +and rep_plus': "n1 <: a1 \ n2 <: a2 \ n1+n2 <: plus' a1 a2" + + +instantiation "fun" :: (type, SL_top) SL_top +begin + +definition "f \ g = (ALL x. f x \ g x)" +definition "f \ g = (\x. f x \ g x)" +definition "Top = (\x. Top)" + +lemma join_apply[simp]: + "(f \ g) x = f x \ g x" +by (simp add: join_fun_def) + +instance +proof + case goal2 thus ?case by (metis le_fun_def preord_class.le_trans) +qed (simp_all add: le_fun_def Top_fun_def) + +end + +subsection "Abstract Interpretation Abstractly" + +text{* Abstract interpretation over abstract values. +Abstract states are simply functions. *} + +locale Abs_Int_Fun = Val_abs +begin + +fun aval' :: "aexp \ (name \ 'a) \ 'a" ("aval\<^isup>#") where +"aval' (N n) _ = num' n" | +"aval' (V x) S = S x" | +"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)" + +abbreviation fun_in_rep (infix "<:" 50) where +"f <: F == ALL x. f x <: F x" + +lemma fun_in_rep_le: "(s::state) <: S \ S \ T \ s <: T" +by (metis le_fun_def le_rep subsetD) + +lemma aval'_sound: "s <: S \ aval a s <: aval' a S" +by (induct a) (auto simp: rep_num' rep_plus') + +fun AI :: "com \ (name \ 'a) \ (name \ 'a)" where +"AI SKIP S = S" | +"AI (x ::= a) S = S(x := aval' a S)" | +"AI (c1;c2) S = AI c2 (AI c1 S)" | +"AI (IF b THEN c1 ELSE c2) S = (AI c1 S) \ (AI c2 S)" | +"AI (WHILE b DO c) S = pfp_above (AI c) S" + +lemma AI_sound: "(c,s) \ t \ s <: S0 \ t <: AI c S0" +proof(induct c arbitrary: s t S0) + case SKIP thus ?case by fastsimp +next + case Assign thus ?case by (auto simp: aval'_sound) +next + case Semi thus ?case by auto +next + case If thus ?case by(auto simp: in_rep_join) +next + case (While b c) + let ?P = "pfp_above (AI c) S0" + have pfp: "AI c ?P \ ?P" and "S0 \ ?P" + by(simp_all add: SL_top_class.pfp_above_pfp) + { fix s t have "(WHILE b DO c,s) \ t \ s <: ?P \ t <: ?P" + proof(induct "WHILE b DO c" s t rule: big_step_induct) + case WhileFalse thus ?case by simp + next + case WhileTrue thus ?case using While.hyps pfp fun_in_rep_le by metis + qed + } + with fun_in_rep_le[OF `s <: S0` `S0 \ ?P`] + show ?case by (metis While(2) AI.simps(5)) +qed + +end + + +text{* Problem: not executable because of the comparison of abstract states, +i.e. functions, in the post-fixedpoint computation. Need to implement +abstract states concretely. *} + + +(* Exercises: show that <= is a preorder if +1. v1 <= v2 = rep v1 <= rep v2 +2. v1 <= v2 = ALL x. lookup v1 x <= lookup v2 x +*) + +end diff -r fe0365331566 -r 17dbd9d9db38 src/HOL/IMP/AbsInt1.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/IMP/AbsInt1.thy Fri Sep 02 19:25:44 2011 +0200 @@ -0,0 +1,219 @@ +(* Author: Tobias Nipkow *) + +theory AbsInt1 +imports AbsInt0_const +begin + +subsection "Backward Analysis of Expressions" + +class L_top_bot = SL_top + +fixes meet :: "'a \ 'a \ 'a" (infixl "\" 65) +and Bot :: "'a" +assumes meet_le1 [simp]: "x \ y \ x" +and meet_le2 [simp]: "x \ y \ y" +and meet_greatest: "x \ y \ x \ z \ x \ y \ z" +assumes bot[simp]: "Bot \ x" + +locale Rep1 = Rep rep for rep :: "'a::L_top_bot \ 'b set" + +assumes inter_rep_subset_rep_meet: "rep a1 \ rep a2 \ rep(a1 \ a2)" +and rep_Bot: "rep Bot = {}" +begin + +lemma in_rep_meet: "x <: a1 \ x <: a2 \ x <: a1 \ a2" +by (metis IntI inter_rep_subset_rep_meet set_mp) + +lemma rep_meet[simp]: "rep(a1 \ a2) = rep a1 \ rep a2" +by (metis equalityI inter_rep_subset_rep_meet le_inf_iff le_rep meet_le1 meet_le2) + +end + + +locale Val_abs1 = Val_abs rep num' plus' + Rep1 rep + for rep :: "'a::L_top_bot \ int set" and num' plus' + +fixes inv_plus' :: "'a \ 'a \ 'a \ 'a * 'a" +and inv_less' :: "'a \ 'a \ bool \ 'a * 'a" +assumes inv_plus': "inv_plus' a1 a2 a = (a1',a2') \ + n1 <: a1 \ n2 <: a2 \ n1+n2 <: a \ n1 <: a1' \ n2 <: a2'" +and inv_less': "inv_less' a1 a2 (n1 + n1 <: a1 \ n2 <: a2 \ n1 <: a1' \ n2 <: a2'" + +datatype 'a up = bot | Up 'a + +instantiation up :: (SL_top)SL_top +begin + +fun le_up where +"Up x \ Up y = (x \ y)" | +"bot \ y = True" | +"Up _ \ bot = False" + +lemma [simp]: "(x \ bot) = (x = bot)" +by (cases x) simp_all + +lemma [simp]: "(Up x \ u) = (EX y. u = Up y & x \ y)" +by (cases u) auto + +fun join_up where +"Up x \ Up y = Up(x \ y)" | +"bot \ y = y" | +"x \ bot = x" + +lemma [simp]: "x \ bot = x" +by (cases x) simp_all + + +definition "Top = Up Top" + +(* register <= as transitive - see orderings *) + +instance proof + case goal1 show ?case by(cases x, simp_all) +next + case goal2 thus ?case + by(cases z, simp, cases y, simp, cases x, auto intro: le_trans) +next + case goal3 thus ?case by(cases x, simp, cases y, simp_all) +next + case goal4 thus ?case by(cases y, simp, cases x, simp_all) +next + case goal5 thus ?case by(cases z, simp, cases y, simp, cases x, simp_all) +next + case goal6 thus ?case by(cases x, simp_all add: Top_up_def) +qed + +end + + +locale Abs_Int1 = Val_abs1 +begin + +(* FIXME avoid duplicating this defn *) +abbreviation astate_in_rep (infix "<:" 50) where +"s <: S == ALL x. s x <: lookup S x" + +abbreviation in_rep_up :: "state \ 'a astate up \ bool" (infix "<::" 50) where +"s <:: S == EX S0. S = Up S0 \ s <: S0" + +lemma in_rep_up_trans: "(s::state) <:: S \ S \ T \ s <:: T" +apply auto +by (metis in_mono le_astate_def le_rep lookup_def top) + +lemma in_rep_join_UpI: "s <:: S1 | s <:: S2 \ s <:: S1 \ S2" +by (metis in_rep_up_trans SL_top_class.join_ge1 SL_top_class.join_ge2) + +fun aval' :: "aexp \ 'a astate up \ 'a" ("aval\<^isup>#") where +"aval' _ bot = Bot" | +"aval' (N n) _ = num' n" | +"aval' (V x) (Up S) = lookup S x" | +"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)" + +lemma aval'_sound: "s <:: S \ aval a s <: aval' a S" +by (induct a) (auto simp: rep_num' rep_plus') + +fun afilter :: "aexp \ 'a \ 'a astate up \ 'a astate up" where +"afilter (N n) a S = (if n <: a then S else bot)" | +"afilter (V x) a S = (case S of bot \ bot | Up S \ + let a' = lookup S x \ a in + if a'\ Bot then bot else Up(update S x a'))" | +"afilter (Plus e1 e2) a S = + (let (a1,a2) = inv_plus' (aval' e1 S) (aval' e2 S) a + in afilter e1 a1 (afilter e2 a2 S))" + +fun bfilter :: "bexp \ bool \ 'a astate up \ 'a astate up" where +"bfilter (B bv) res S = (if bv=res then S else bot)" | +"bfilter (Not b) res S = bfilter b (\ res) S" | +"bfilter (And b1 b2) res S = + (if res then bfilter b1 True (bfilter b2 True S) + else bfilter b1 False S \ bfilter b2 False S)" | +"bfilter (Less e1 e2) res S = + (let (res1,res2) = inv_less' (aval' e1 S) (aval' e2 S) res + in afilter e1 res1 (afilter e2 res2 S))" + +lemma afilter_sound: "s <:: S \ aval e s <: a \ s <:: afilter e a S" +proof(induct e arbitrary: a S) + case N thus ?case by simp +next + case (V x) + obtain S' where "S = Up S'" and "s <: S'" using `s <:: S` by auto + moreover hence "s x <: lookup S' x" by(simp) + moreover have "s x <: a" using V by simp + ultimately show ?case using V(1) + by(simp add: lookup_update Let_def) + (metis le_rep emptyE in_rep_meet rep_Bot subset_empty) +next + case (Plus e1 e2) thus ?case + using inv_plus'[OF _ aval'_sound[OF Plus(3)] aval'_sound[OF Plus(3)]] + by (auto split: prod.split) +qed + +lemma bfilter_sound: "s <:: S \ bv = bval b s \ s <:: bfilter b bv S" +proof(induct b arbitrary: S bv) + case B thus ?case by simp +next + case (Not b) thus ?case by simp +next + case (And b1 b2) thus ?case by (auto simp: in_rep_join_UpI) +next + case (Less e1 e2) thus ?case + by (auto split: prod.split) + (metis afilter_sound inv_less' aval'_sound Less) +qed + +fun AI :: "com \ 'a astate up \ 'a astate up" where +"AI SKIP S = S" | +"AI (x ::= a) S = + (case S of bot \ bot | Up S \ Up(update S x (aval' a (Up S))))" | +"AI (c1;c2) S = AI c2 (AI c1 S)" | +"AI (IF b THEN c1 ELSE c2) S = + AI c1 (bfilter b True S) \ AI c2 (bfilter b False S)" | +"AI (WHILE b DO c) S = + bfilter b False (pfp_above (%S. AI c (bfilter b True S)) S)" + +lemma AI_sound: "(c,s) \ t \ s <:: S \ t <:: AI c S" +proof(induct c arbitrary: s t S) + case SKIP thus ?case by fastsimp +next + case Assign thus ?case + by (auto simp: lookup_update aval'_sound) +next + case Semi thus ?case by fastsimp +next + case If thus ?case by (auto simp: in_rep_join_UpI bfilter_sound) +next + case (While b c) + let ?P = "pfp_above (%S. AI c (bfilter b True S)) S" + have pfp: "AI c (bfilter b True ?P) \ ?P" and "S \ ?P" + by (rule pfp_above_pfp(1), rule pfp_above_pfp(2)) + { fix s t + have "(WHILE b DO c,s) \ t \ s <:: ?P \ + t <:: bfilter b False ?P" + proof(induct "WHILE b DO c" s t rule: big_step_induct) + case WhileFalse thus ?case by(metis bfilter_sound) + next + case WhileTrue show ?case + by(rule WhileTrue, rule in_rep_up_trans[OF _ pfp], + rule While.hyps[OF WhileTrue(2)], + rule bfilter_sound[OF WhileTrue.prems], simp add: WhileTrue(1)) + qed + } + with in_rep_up_trans[OF `s <:: S` `S \ ?P`] While(2,3) AI.simps(5) + show ?case by simp +qed + +text{* Exercise: *} + +lemma afilter_strict: "afilter e res bot = bot" +by (induct e arbitrary: res) simp_all + +lemma bfilter_strict: "bfilter b res bot = bot" +by (induct b arbitrary: res) (simp_all add: afilter_strict) + +lemma pfp_strict: "f bot = bot \ pfp_above f bot = bot" +by(simp add: pfp_above_def pfp_def eval_nat_numeral) + +lemma AI_strict: "AI c bot = bot" +by(induct c) (simp_all add: bfilter_strict pfp_strict) + +end + +end diff -r fe0365331566 -r 17dbd9d9db38 src/HOL/IMP/AbsInt1_ivl.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/IMP/AbsInt1_ivl.thy Fri Sep 02 19:25:44 2011 +0200 @@ -0,0 +1,231 @@ +(* Author: Tobias Nipkow *) + +theory AbsInt1_ivl +imports AbsInt1 +begin + +subsection "Interval Analysis" + +datatype ivl = I "int option" "int option" + +definition "rep_ivl i = + (case i of I (Some l) (Some h) \ {l..h} | I (Some l) None \ {l..} + | I None (Some h) \ {..h} | I None None \ UNIV)" + +instantiation option :: (plus)plus +begin + +fun plus_option where +"Some x + Some y = Some(x+y)" | +"_ + _ = None" + +instance proof qed + +end + +definition empty where "empty = I (Some 1) (Some 0)" + +fun is_empty where +"is_empty(I (Some l) (Some h)) = (h (case h of Some h \ h False) | None \ False)" +by(auto split:option.split) + +lemma [simp]: "is_empty i \ rep_ivl i = {}" +by(auto simp add: rep_ivl_def split: ivl.split option.split) + +definition "plus_ivl i1 i2 = (if is_empty i1 | is_empty i2 then empty + else case (i1,i2) of (I l1 h1, I l2 h2) \ I (l1+l2) (h1+h2))" + +instantiation ivl :: SL_top +begin + +definition le_option :: "bool \ int option \ int option \ bool" where +"le_option pos x y = + (case x of (Some i) \ (case y of (Some j) \ (i<=j) | None \ pos) + | None \ (case y of Some j \ (~pos) | None \ True))" + +fun le_aux where +"le_aux (I l1 h1) (I l2 h2) = (le_option False l2 l1 & le_option True h1 h2)" + +definition le_ivl where +"i1 \ i2 = + (if is_empty i1 then True else + if is_empty i2 then False else le_aux i1 i2)" + +definition min_option :: "bool \ int option \ int option \ int option" where +"min_option pos o1 o2 = (if le_option pos o1 o2 then o1 else o2)" + +definition max_option :: "bool \ int option \ int option \ int option" where +"max_option pos o1 o2 = (if le_option pos o1 o2 then o2 else o1)" + +definition "i1 \ i2 = + (if is_empty i1 then i2 else if is_empty i2 then i1 + else case (i1,i2) of (I l1 h1, I l2 h2) \ + I (min_option False l1 l2) (max_option True h1 h2))" + +definition "Top = I None None" + +instance +proof + case goal1 thus ?case + by(cases x, simp add: le_ivl_def le_option_def split: option.split) +next + case goal2 thus ?case + by(cases x, cases y, cases z, auto simp: le_ivl_def le_option_def split: option.splits if_splits) +next + case goal3 thus ?case + by(cases x, cases y, simp add: le_ivl_def join_ivl_def le_option_def min_option_def max_option_def split: option.splits) +next + case goal4 thus ?case + by(cases x, cases y, simp add: le_ivl_def join_ivl_def le_option_def min_option_def max_option_def split: option.splits) +next + case goal5 thus ?case + by(cases x, cases y, cases z, auto simp add: le_ivl_def join_ivl_def le_option_def min_option_def max_option_def split: option.splits if_splits) +next + case goal6 thus ?case + by(cases x, simp add: Top_ivl_def le_ivl_def le_option_def split: option.split) +qed + +end + + +instantiation ivl :: L_top_bot +begin + +definition "i1 \ i2 = (if is_empty i1 | is_empty i2 then empty + else case (i1,i2) of (I l1 h1, I l2 h2) \ I (max_option False l1 l2) (min_option True h1 h2))" + +definition "Bot = empty" + +instance +proof + case goal1 thus ?case + by (simp add:meet_ivl_def empty_def meet_ivl_def le_ivl_def le_option_def max_option_def min_option_def split: ivl.splits option.splits) +next + case goal2 thus ?case + by (simp add:meet_ivl_def empty_def meet_ivl_def le_ivl_def le_option_def max_option_def min_option_def split: ivl.splits option.splits) +next + case goal3 thus ?case + by (cases x, cases y, cases z, auto simp add: le_ivl_def meet_ivl_def empty_def le_option_def max_option_def min_option_def split: option.splits if_splits) +next + case goal4 show ?case by(cases x, simp add: Bot_ivl_def empty_def le_ivl_def) +qed + +end + +instantiation option :: (minus)minus +begin + +fun minus_option where +"Some x - Some y = Some(x-y)" | +"_ - _ = None" + +instance proof qed + +end + +definition "minus_ivl i1 i2 = + (if is_empty i1 | is_empty i2 then empty + else case (i1,i2) of + (I l1 h1, I l2 h2) \ I (l1-h2) (h1-l2))" + +lemma rep_minus_ivl: + "n1 : rep_ivl i1 \ n2 : rep_ivl i2 \ n1-n2 : rep_ivl(minus_ivl i1 i2)" +by(auto simp add: minus_ivl_def rep_ivl_def split: ivl.splits option.splits) + + +definition "inv_plus_ivl i1 i2 i = + (if is_empty i then empty + else i1 \ minus_ivl i i2, i2 \ minus_ivl i i1)" + +fun inv_less_ivl :: "ivl \ ivl \ bool \ ivl * ivl" where +"inv_less_ivl (I l1 h1) (I l2 h2) res = + (if res + then (I l1 (min_option True h1 (h2 - Some 1)), I (max_option False (l1 + Some 1) l2) h2) + else (I (max_option False l1 l2) h1, I l2 (min_option True h1 h2)))" + +interpretation Rep rep_ivl +proof + case goal1 thus ?case + by(auto simp: rep_ivl_def le_ivl_def le_option_def split: ivl.split option.split if_splits) +qed + +interpretation Val_abs rep_ivl "%n. I (Some n) (Some n)" plus_ivl +proof + case goal1 thus ?case by(simp add: rep_ivl_def) +next + case goal2 thus ?case + by(auto simp add: rep_ivl_def plus_ivl_def split: ivl.split option.splits) +qed + +interpretation Rep1 rep_ivl +proof + case goal1 thus ?case + by(auto simp add: rep_ivl_def meet_ivl_def empty_def min_option_def max_option_def split: ivl.split option.split) +next + case goal2 show ?case by(auto simp add: Bot_ivl_def rep_ivl_def empty_def) +qed + +interpretation Val_abs1 rep_ivl "%n. I (Some n) (Some n)" plus_ivl inv_plus_ivl inv_less_ivl +proof + case goal1 thus ?case + by(auto simp add: inv_plus_ivl_def) + (metis rep_minus_ivl add_diff_cancel add_commute)+ +next + case goal2 thus ?case + by(cases a1, cases a2, + auto simp: rep_ivl_def min_option_def max_option_def le_option_def split: if_splits option.splits) +qed + +interpretation Abs_Int1 rep_ivl "%n. I (Some n) (Some n)" plus_ivl inv_plus_ivl inv_less_ivl +defines afilter_ivl is afilter +and bfilter_ivl is bfilter +and AI_ivl is AI +and aval_ivl is aval' +.. + + +fun list_up where +"list_up bot = bot" | +"list_up (Up x) = Up(list x)" + +value [code] "list_up(afilter_ivl (N 5) (I (Some 4) (Some 5)) Top)" +value [code] "list_up(afilter_ivl (N 5) (I (Some 4) (Some 4)) Top)" +value [code] "list_up(afilter_ivl (V ''x'') (I (Some 4) (Some 4)) + (Up(FunDom(Top(''x'':=I (Some 5) (Some 6))) [''x''])))" +value [code] "list_up(afilter_ivl (V ''x'') (I (Some 4) (Some 5)) + (Up(FunDom(Top(''x'':=I (Some 5) (Some 6))) [''x''])))" +value [code] "list_up(afilter_ivl (Plus (V ''x'') (V ''x'')) (I (Some 0) (Some 10)) + (Up(FunDom(Top(''x'':= I (Some 0) (Some 100)))[''x''])))" +value [code] "list_up(afilter_ivl (Plus (V ''x'') (N 7)) (I (Some 0) (Some 10)) + (Up(FunDom(Top(''x'':= I (Some 0) (Some 100)))[''x''])))" + +value [code] "list_up(bfilter_ivl (Less (V ''x'') (V ''x'')) True + (Up(FunDom(Top(''x'':= I (Some 0) (Some 0)))[''x''])))" +value [code] "list_up(bfilter_ivl (Less (V ''x'') (V ''x'')) True + (Up(FunDom(Top(''x'':= I (Some 0) (Some 2)))[''x''])))" +value [code] "list_up(bfilter_ivl (Less (V ''x'') (Plus (N 10) (V ''y''))) True + (Up(FunDom(Top(''x'':= I (Some 15) (Some 20),''y'':= I (Some 5) (Some 7)))[''x'', ''y''])))" + +definition "test_ivl1 = + ''y'' ::= N 7; + IF Less (V ''x'') (V ''y'') + THEN ''y'' ::= Plus (V ''y'') (V ''x'') + ELSE ''x'' ::= Plus (V ''x'') (V ''y'')" +value [code] "list_up(AI_ivl test_ivl1 Top)" + +value "list_up (AI_ivl test3_const Top)" + +value "list_up (AI_ivl test5_const Top)" + +value "list_up (AI_ivl test6_const Top)" + +definition "test2_ivl = + ''y'' ::= N 7; + WHILE Less (V ''x'') (N 100) DO ''y'' ::= Plus (V ''y'') (N 1)" +value [code] "list_up(AI_ivl test2_ivl Top)" + +end diff -r fe0365331566 -r 17dbd9d9db38 src/HOL/IMP/Astate.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/IMP/Astate.thy Fri Sep 02 19:25:44 2011 +0200 @@ -0,0 +1,46 @@ +(* Author: Tobias Nipkow *) + +theory Astate +imports AbsInt0_fun +begin + +subsection "Abstract State with Computable Ordering" + +text{* A concrete type of state with computable @{text"\"}: *} + +datatype 'a astate = FunDom "string \ 'a" "string list" + +fun "fun" where "fun (FunDom f _) = f" +fun dom where "dom (FunDom _ A) = A" + +definition "list S = [(x,fun S x). x \ dom S]" + +definition "lookup F x = (if x : set(dom F) then fun F x else Top)" + +definition "update F x y = + FunDom ((fun F)(x:=y)) (if x : set(dom F) then dom F else x # dom F)" + +lemma lookup_update: "lookup (update S x y) = (lookup S)(x:=y)" +by(rule ext)(auto simp: lookup_def update_def) + +instantiation astate :: (SL_top) SL_top +begin + +definition "le_astate F G = (ALL x : set(dom G). lookup F x \ fun G x)" + +definition +"join_astate F G = + FunDom (%x. fun F x \ fun G x) ([x\dom F. x : set(dom G)])" + +definition "Top = FunDom (%x. Top) []" + +instance +proof + case goal2 thus ?case + apply(auto simp: le_astate_def) + by (metis lookup_def preord_class.le_trans top) +qed (auto simp: le_astate_def lookup_def join_astate_def Top_astate_def) + +end + +end diff -r fe0365331566 -r 17dbd9d9db38 src/HOL/IMP/ROOT.ML --- a/src/HOL/IMP/ROOT.ML Fri Sep 02 18:17:45 2011 +0200 +++ b/src/HOL/IMP/ROOT.ML Fri Sep 02 19:25:44 2011 +0200 @@ -1,3 +1,5 @@ +no_document use_thys ["~~/src/HOL/ex/Interpretation_with_Defs"]; + use_thys ["BExp", "ASM", @@ -10,6 +12,7 @@ "Def_Ass_Sound_Big", "Def_Ass_Sound_Small", "Live", + "AbsInt1_ivl", "Hoare_Examples", "VC", "HoareT", diff -r fe0365331566 -r 17dbd9d9db38 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Fri Sep 02 18:17:45 2011 +0200 +++ b/src/HOL/IsaMakefile Fri Sep 02 19:25:44 2011 +0200 @@ -511,7 +511,9 @@ HOL-IMP: HOL $(OUT)/HOL-IMP -$(OUT)/HOL-IMP: $(OUT)/HOL IMP/ASM.thy IMP/AExp.thy IMP/BExp.thy \ +$(OUT)/HOL-IMP: $(OUT)/HOL IMP/AbsInt0_fun.thy IMP/Astate.thy \ + IMP/AbsInt0.thy IMP/AbsInt0_const.thy IMP/AbsInt1.thy IMP/AbsInt1_ivl.thy \ + IMP/ASM.thy IMP/AExp.thy IMP/BExp.thy \ IMP/Big_Step.thy IMP/C_like.thy IMP/Com.thy IMP/Compiler.thy \ IMP/Comp_Rev.thy IMP/Def_Ass.thy IMP/Def_Ass_Big.thy IMP/Def_Ass_Exp.thy \ IMP/Def_Ass_Small.thy IMP/Def_Ass_Sound_Big.thy \