# HG changeset patch # User nipkow # Date 1479914922 -3600 # Node ID 49b78f1f9e016af29601151edc3bfddce1ba1e75 # Parent e44f5c123f2667787c9df0694aabca060249665a moved IMP/Abs_Int_ITP to AFP/Abs_Int_ITP2012 diff -r e44f5c123f26 -r 49b78f1f9e01 src/HOL/IMP/Abs_Int_ITP/ACom_ITP.thy --- a/src/HOL/IMP/Abs_Int_ITP/ACom_ITP.thy Tue Nov 22 18:36:59 2016 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,125 +0,0 @@ -(* Author: Tobias Nipkow *) - -theory ACom_ITP -imports "../Com" -begin - -subsection "Annotated Commands" - -datatype 'a acom = - SKIP 'a ("SKIP {_}" 61) | - Assign vname aexp 'a ("(_ ::= _/ {_})" [1000, 61, 0] 61) | - Seq "('a acom)" "('a acom)" ("_;;//_" [60, 61] 60) | - If bexp "('a acom)" "('a acom)" 'a - ("(IF _/ THEN _/ ELSE _//{_})" [0, 0, 61, 0] 61) | - While 'a bexp "('a acom)" 'a - ("({_}//WHILE _/ DO (_)//{_})" [0, 0, 61, 0] 61) - -fun post :: "'a acom \'a" where -"post (SKIP {P}) = P" | -"post (x ::= e {P}) = P" | -"post (c1;; c2) = post c2" | -"post (IF b THEN c1 ELSE c2 {P}) = P" | -"post ({Inv} WHILE b DO c {P}) = P" - -fun strip :: "'a acom \ com" where -"strip (SKIP {P}) = com.SKIP" | -"strip (x ::= e {P}) = (x ::= e)" | -"strip (c1;;c2) = (strip c1;; strip c2)" | -"strip (IF b THEN c1 ELSE c2 {P}) = (IF b THEN strip c1 ELSE strip c2)" | -"strip ({Inv} WHILE b DO c {P}) = (WHILE b DO strip c)" - -fun anno :: "'a \ com \ 'a acom" where -"anno a com.SKIP = SKIP {a}" | -"anno a (x ::= e) = (x ::= e {a})" | -"anno a (c1;;c2) = (anno a c1;; anno a c2)" | -"anno a (IF b THEN c1 ELSE c2) = - (IF b THEN anno a c1 ELSE anno a c2 {a})" | -"anno a (WHILE b DO c) = - ({a} WHILE b DO anno a c {a})" - -fun annos :: "'a acom \ 'a list" where -"annos (SKIP {a}) = [a]" | -"annos (x::=e {a}) = [a]" | -"annos (C1;;C2) = annos C1 @ annos C2" | -"annos (IF b THEN C1 ELSE C2 {a}) = a # annos C1 @ annos C2" | -"annos ({i} WHILE b DO C {a}) = i # a # annos C" - -fun map_acom :: "('a \ 'b) \ 'a acom \ 'b acom" where -"map_acom f (SKIP {P}) = SKIP {f P}" | -"map_acom f (x ::= e {P}) = (x ::= e {f P})" | -"map_acom f (c1;;c2) = (map_acom f c1;; map_acom f c2)" | -"map_acom f (IF b THEN c1 ELSE c2 {P}) = - (IF b THEN map_acom f c1 ELSE map_acom f c2 {f P})" | -"map_acom f ({Inv} WHILE b DO c {P}) = - ({f Inv} WHILE b DO map_acom f c {f P})" - - -lemma post_map_acom[simp]: "post(map_acom f c) = f(post c)" -by (induction c) simp_all - -lemma strip_acom[simp]: "strip (map_acom f c) = strip c" -by (induction c) auto - -lemma map_acom_SKIP: - "map_acom f c = SKIP {S'} \ (\S. c = SKIP {S} \ S' = f S)" -by (cases c) auto - -lemma map_acom_Assign: - "map_acom f c = x ::= e {S'} \ (\S. c = x::=e {S} \ S' = f S)" -by (cases c) auto - -lemma map_acom_Seq: - "map_acom f c = c1';;c2' \ - (\c1 c2. c = c1;;c2 \ map_acom f c1 = c1' \ map_acom f c2 = c2')" -by (cases c) auto - -lemma map_acom_If: - "map_acom f c = IF b THEN c1' ELSE c2' {S'} \ - (\S c1 c2. c = IF b THEN c1 ELSE c2 {S} \ map_acom f c1 = c1' \ map_acom f c2 = c2' \ S' = f S)" -by (cases c) auto - -lemma map_acom_While: - "map_acom f w = {I'} WHILE b DO c' {P'} \ - (\I P c. w = {I} WHILE b DO c {P} \ map_acom f c = c' \ I' = f I \ P' = f P)" -by (cases w) auto - - -lemma strip_anno[simp]: "strip (anno a c) = c" -by(induct c) simp_all - -lemma strip_eq_SKIP: - "strip c = com.SKIP \ (EX P. c = SKIP {P})" -by (cases c) simp_all - -lemma strip_eq_Assign: - "strip c = x::=e \ (EX P. c = x::=e {P})" -by (cases c) simp_all - -lemma strip_eq_Seq: - "strip c = c1;;c2 \ (EX d1 d2. c = d1;;d2 & strip d1 = c1 & strip d2 = c2)" -by (cases c) simp_all - -lemma strip_eq_If: - "strip c = IF b THEN c1 ELSE c2 \ - (EX d1 d2 P. c = IF b THEN d1 ELSE d2 {P} & strip d1 = c1 & strip d2 = c2)" -by (cases c) simp_all - -lemma strip_eq_While: - "strip c = WHILE b DO c1 \ - (EX I d1 P. c = {I} WHILE b DO d1 {P} & strip d1 = c1)" -by (cases c) simp_all - - -lemma set_annos_anno[simp]: "set (annos (anno a C)) = {a}" -by(induction C)(auto) - -lemma size_annos_same: "strip C1 = strip C2 \ size(annos C1) = size(annos C2)" -apply(induct C2 arbitrary: C1) -apply (auto simp: strip_eq_SKIP strip_eq_Assign strip_eq_Seq strip_eq_If strip_eq_While) -done - -lemmas size_annos_same2 = eqTrueI[OF size_annos_same] - - -end diff -r e44f5c123f26 -r 49b78f1f9e01 src/HOL/IMP/Abs_Int_ITP/Abs_Int0_ITP.thy --- a/src/HOL/IMP/Abs_Int_ITP/Abs_Int0_ITP.thy Tue Nov 22 18:36:59 2016 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,397 +0,0 @@ -(* Author: Tobias Nipkow *) - -theory Abs_Int0_ITP -imports - "~~/src/HOL/Library/While_Combinator" - Collecting_ITP -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" -begin - -definition mono where "mono f = (\x y. x \ y \ f x \ f y)" - -lemma monoD: "mono f \ x \ y \ f x \ f y" by(simp add: mono_def) - -lemma mono_comp: "mono f \ mono g \ mono (g o f)" -by(simp add: mono_def) - -declare le_trans[trans] - -end - -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 \ \" -begin - -lemma join_le_iff[simp]: "x \ y \ z \ x \ z \ y \ z" -by (metis join_ge1 join_ge2 join_least le_trans) - -lemma le_join_disj: "x \ y \ x \ z \ x \ y \ z" -by (metis join_ge1 join_ge2 le_trans) - -end - -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 "\ = (\x. \)" - -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 - - -instantiation acom :: (preord) preord -begin - -fun le_acom :: "('a::preord)acom \ 'a acom \ bool" where -"le_acom (SKIP {S}) (SKIP {S'}) = (S \ S')" | -"le_acom (x ::= e {S}) (x' ::= e' {S'}) = (x=x' \ e=e' \ S \ S')" | -"le_acom (c1;;c2) (c1';;c2') = (le_acom c1 c1' \ le_acom c2 c2')" | -"le_acom (IF b THEN c1 ELSE c2 {S}) (IF b' THEN c1' ELSE c2' {S'}) = - (b=b' \ le_acom c1 c1' \ le_acom c2 c2' \ S \ S')" | -"le_acom ({Inv} WHILE b DO c {P}) ({Inv'} WHILE b' DO c' {P'}) = - (b=b' \ le_acom c c' \ Inv \ Inv' \ P \ P')" | -"le_acom _ _ = False" - -lemma [simp]: "SKIP {S} \ c \ (\S'. c = SKIP {S'} \ S \ S')" -by (cases c) auto - -lemma [simp]: "x ::= e {S} \ c \ (\S'. c = x ::= e {S'} \ S \ S')" -by (cases c) auto - -lemma [simp]: "c1;;c2 \ c \ (\c1' c2'. c = c1';;c2' \ c1 \ c1' \ c2 \ c2')" -by (cases c) auto - -lemma [simp]: "IF b THEN c1 ELSE c2 {S} \ c \ - (\c1' c2' S'. c = IF b THEN c1' ELSE c2' {S'} \ c1 \ c1' \ c2 \ c2' \ S \ S')" -by (cases c) auto - -lemma [simp]: "{Inv} WHILE b DO c {P} \ w \ - (\Inv' c' P'. w = {Inv'} WHILE b DO c' {P'} \ c \ c' \ Inv \ Inv' \ P \ P')" -by (cases w) auto - -instance -proof - case goal1 thus ?case by (induct x) auto -next - case goal2 thus ?case - apply(induct x y arbitrary: z rule: le_acom.induct) - apply (auto intro: le_trans) - done -qed - -end - - -subsubsection "Lifting" - -instantiation option :: (preord)preord -begin - -fun le_option where -"Some x \ Some y = (x \ y)" | -"None \ y = True" | -"Some _ \ None = False" - -lemma [simp]: "(x \ None) = (x = None)" -by (cases x) simp_all - -lemma [simp]: "(Some x \ u) = (\y. u = Some y & x \ y)" -by (cases u) auto - -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) -qed - -end - -instantiation option :: (SL_top)SL_top -begin - -fun join_option where -"Some x \ Some y = Some(x \ y)" | -"None \ y = y" | -"x \ None = x" - -lemma join_None2[simp]: "x \ None = x" -by (cases x) simp_all - -definition "\ = Some \" - -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) -qed - -end - -definition bot_acom :: "com \ ('a::SL_top)option acom" ("\\<^sub>c") where -"\\<^sub>c = anno None" - -lemma strip_bot_acom[simp]: "strip(\\<^sub>c c) = c" -by(simp add: bot_acom_def) - -lemma bot_acom[rule_format]: "strip c' = c \ \\<^sub>c c \ c'" -apply(induct c arbitrary: c') -apply (simp_all add: bot_acom_def) - apply(induct_tac c') - apply simp_all - apply(induct_tac c') - apply simp_all - apply(induct_tac c') - apply simp_all - apply(induct_tac c') - apply simp_all - apply(induct_tac c') - apply simp_all -done - - -subsubsection "Post-fixed point iteration" - -definition - pfp :: "(('a::preord) \ 'a) \ 'a \ 'a option" where -"pfp f = while_option (\x. \ f x \ x) f" - -lemma pfp_pfp: assumes "pfp f x0 = Some x" shows "f x \ x" -using while_option_stop[OF assms[simplified pfp_def]] by simp - -lemma pfp_least: -assumes mono: "\x y. x \ y \ f x \ f y" -and "f p \ p" and "x0 \ p" and "pfp f x0 = Some x" shows "x \ p" -proof- - { fix x assume "x \ p" - hence "f x \ f p" by(rule mono) - from this `f p \ p` have "f x \ p" by(rule le_trans) - } - thus "x \ p" using assms(2-) while_option_rule[where P = "%x. x \ p"] - unfolding pfp_def by blast -qed - -definition - lpfp\<^sub>c :: "(('a::SL_top)option acom \ 'a option acom) \ com \ 'a option acom option" where -"lpfp\<^sub>c f c = pfp f (\\<^sub>c c)" - -lemma lpfpc_pfp: "lpfp\<^sub>c f c0 = Some c \ f c \ c" -by(simp add: pfp_pfp lpfp\<^sub>c_def) - -lemma strip_pfp: -assumes "\x. g(f x) = g x" and "pfp f x0 = Some x" shows "g x = g x0" -using assms while_option_rule[where P = "%x. g x = g x0" and c = f] -unfolding pfp_def by metis - -lemma strip_lpfpc: assumes "\c. strip(f c) = strip c" and "lpfp\<^sub>c f c = Some c'" -shows "strip c' = c" -using assms(1) strip_pfp[OF _ assms(2)[simplified lpfp\<^sub>c_def]] -by(metis strip_bot_acom) - -lemma lpfpc_least: -assumes mono: "\x y. x \ y \ f x \ f y" -and "strip p = c0" and "f p \ p" and lp: "lpfp\<^sub>c f c0 = Some c" shows "c \ p" -using pfp_least[OF _ _ bot_acom[OF `strip p = c0`] lp[simplified lpfp\<^sub>c_def]] - mono `f p \ p` -by blast - - -subsection "Abstract Interpretation" - -definition \_fun :: "('a \ 'b set) \ ('c \ 'a) \ ('c \ 'b)set" where -"\_fun \ F = {f. \x. f x \ \(F x)}" - -fun \_option :: "('a \ 'b set) \ 'a option \ 'b set" where -"\_option \ None = {}" | -"\_option \ (Some a) = \ a" - -text{* The interface for abstract values: *} - -locale Val_abs = -fixes \ :: "'av::SL_top \ val set" - assumes mono_gamma: "a \ b \ \ a \ \ b" - and gamma_Top[simp]: "\ \ = UNIV" -fixes num' :: "val \ 'av" -and plus' :: "'av \ 'av \ 'av" - assumes gamma_num': "n : \(num' n)" - and gamma_plus': - "n1 : \ a1 \ n2 : \ a2 \ n1+n2 : \(plus' a1 a2)" - -type_synonym 'av st = "(vname \ 'av)" - -locale Abs_Int_Fun = Val_abs \ for \ :: "'av::SL_top \ val set" -begin - -fun aval' :: "aexp \ 'av st \ 'av" where -"aval' (N n) S = num' n" | -"aval' (V x) S = S x" | -"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)" - -fun step' :: "'av st option \ 'av st option acom \ 'av st option acom" - where -"step' S (SKIP {P}) = (SKIP {S})" | -"step' S (x ::= e {P}) = - x ::= e {case S of None \ None | Some S \ Some(S(x := aval' e S))}" | -"step' S (c1;; c2) = step' S c1;; step' (post c1) c2" | -"step' S (IF b THEN c1 ELSE c2 {P}) = - IF b THEN step' S c1 ELSE step' S c2 {post c1 \ post c2}" | -"step' S ({Inv} WHILE b DO c {P}) = - {S \ post c} WHILE b DO (step' Inv c) {Inv}" - -definition AI :: "com \ 'av st option acom option" where -"AI = lpfp\<^sub>c (step' \)" - - -lemma strip_step'[simp]: "strip(step' S c) = strip c" -by(induct c arbitrary: S) (simp_all add: Let_def) - - -abbreviation \\<^sub>f :: "'av st \ state set" -where "\\<^sub>f == \_fun \" - -abbreviation \\<^sub>o :: "'av st option \ state set" -where "\\<^sub>o == \_option \\<^sub>f" - -abbreviation \\<^sub>c :: "'av st option acom \ state set acom" -where "\\<^sub>c == map_acom \\<^sub>o" - -lemma gamma_f_Top[simp]: "\\<^sub>f Top = UNIV" -by(simp add: Top_fun_def \_fun_def) - -lemma gamma_o_Top[simp]: "\\<^sub>o Top = UNIV" -by (simp add: Top_option_def) - -(* FIXME (maybe also le \ sqle?) *) - -lemma mono_gamma_f: "f \ g \ \\<^sub>f f \ \\<^sub>f g" -by(auto simp: le_fun_def \_fun_def dest: mono_gamma) - -lemma mono_gamma_o: - "sa \ sa' \ \\<^sub>o sa \ \\<^sub>o sa'" -by(induction sa sa' rule: le_option.induct)(simp_all add: mono_gamma_f) - -lemma mono_gamma_c: "ca \ ca' \ \\<^sub>c ca \ \\<^sub>c ca'" -by (induction ca ca' rule: le_acom.induct) (simp_all add:mono_gamma_o) - -text{* Soundness: *} - -lemma aval'_sound: "s : \\<^sub>f S \ aval a s : \(aval' a S)" -by (induct a) (auto simp: gamma_num' gamma_plus' \_fun_def) - -lemma in_gamma_update: - "\ s : \\<^sub>f S; i : \ a \ \ s(x := i) : \\<^sub>f(S(x := a))" -by(simp add: \_fun_def) - -lemma step_preserves_le: - "\ S \ \\<^sub>o S'; c \ \\<^sub>c c' \ \ step S c \ \\<^sub>c (step' S' c')" -proof(induction c arbitrary: c' S S') - case SKIP thus ?case by(auto simp:SKIP_le map_acom_SKIP) -next - case Assign thus ?case - by (fastforce simp: Assign_le map_acom_Assign intro: aval'_sound in_gamma_update - split: option.splits del:subsetD) -next - case Seq thus ?case apply (auto simp: Seq_le map_acom_Seq) - by (metis le_post post_map_acom) -next - case (If b c1 c2 P) - then obtain c1' c2' P' where - "c' = IF b THEN c1' ELSE c2' {P'}" - "P \ \\<^sub>o P'" "c1 \ \\<^sub>c c1'" "c2 \ \\<^sub>c c2'" - by (fastforce simp: If_le map_acom_If) - moreover have "post c1 \ \\<^sub>o(post c1' \ post c2')" - by (metis (no_types) `c1 \ \\<^sub>c c1'` join_ge1 le_post mono_gamma_o order_trans post_map_acom) - moreover have "post c2 \ \\<^sub>o(post c1' \ post c2')" - by (metis (no_types) `c2 \ \\<^sub>c c2'` join_ge2 le_post mono_gamma_o order_trans post_map_acom) - ultimately show ?case using `S \ \\<^sub>o S'` by (simp add: If.IH subset_iff) -next - case (While I b c1 P) - then obtain c1' I' P' where - "c' = {I'} WHILE b DO c1' {P'}" - "I \ \\<^sub>o I'" "P \ \\<^sub>o P'" "c1 \ \\<^sub>c c1'" - by (fastforce simp: map_acom_While While_le) - moreover have "S \ post c1 \ \\<^sub>o (S' \ post c1')" - using `S \ \\<^sub>o S'` le_post[OF `c1 \ \\<^sub>c c1'`, simplified] - by (metis (no_types) join_ge1 join_ge2 le_sup_iff mono_gamma_o order_trans) - ultimately show ?case by (simp add: While.IH subset_iff) -qed - -lemma AI_sound: "AI c = Some c' \ CS c \ \\<^sub>c c'" -proof(simp add: CS_def AI_def) - assume 1: "lpfp\<^sub>c (step' \) c = Some c'" - have 2: "step' \ c' \ c'" by(rule lpfpc_pfp[OF 1]) - have 3: "strip (\\<^sub>c (step' \ c')) = c" - by(simp add: strip_lpfpc[OF _ 1]) - have "lfp (step UNIV) c \ \\<^sub>c (step' \ c')" - proof(rule lfp_lowerbound[simplified,OF 3]) - show "step UNIV (\\<^sub>c (step' \ c')) \ \\<^sub>c (step' \ c')" - proof(rule step_preserves_le[OF _ _]) - show "UNIV \ \\<^sub>o \" by simp - show "\\<^sub>c (step' \ c') \ \\<^sub>c c'" by(rule mono_gamma_c[OF 2]) - qed - qed - with 2 show "lfp (step UNIV) c \ \\<^sub>c c'" - by (blast intro: mono_gamma_c order_trans) -qed - -end - - -subsubsection "Monotonicity" - -lemma mono_post: "c \ c' \ post c \ post c'" -by(induction c c' rule: le_acom.induct) (auto) - -locale Abs_Int_Fun_mono = Abs_Int_Fun + -assumes mono_plus': "a1 \ b1 \ a2 \ b2 \ plus' a1 a2 \ plus' b1 b2" -begin - -lemma mono_aval': "S \ S' \ aval' e S \ aval' e S'" -by(induction e)(auto simp: le_fun_def mono_plus') - -lemma mono_update: "a \ a' \ S \ S' \ S(x := a) \ S'(x := a')" -by(simp add: le_fun_def) - -lemma mono_step': "S \ S' \ c \ c' \ step' S c \ step' S' c'" -apply(induction c c' arbitrary: S S' rule: le_acom.induct) -apply (auto simp: Let_def mono_update mono_aval' mono_post le_join_disj - split: option.split) -done - -end - -text{* Problem: not executable because of the comparison of abstract states, -i.e. functions, in the post-fixedpoint computation. *} - -end diff -r e44f5c123f26 -r 49b78f1f9e01 src/HOL/IMP/Abs_Int_ITP/Abs_Int1_ITP.thy --- a/src/HOL/IMP/Abs_Int_ITP/Abs_Int1_ITP.thy Tue Nov 22 18:36:59 2016 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,372 +0,0 @@ -(* Author: Tobias Nipkow *) - -theory Abs_Int1_ITP -imports Abs_State_ITP -begin - -subsection "Computable Abstract Interpretation" - -text{* Abstract interpretation over type @{text st} instead of -functions. *} - -context Gamma -begin - -fun aval' :: "aexp \ 'av st \ 'av" where -"aval' (N n) S = num' n" | -"aval' (V x) S = lookup S x" | -"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)" - -lemma aval'_sound: "s : \\<^sub>f S \ aval a s : \(aval' a S)" -by (induction a) (auto simp: gamma_num' gamma_plus' \_st_def lookup_def) - -end - -text{* The for-clause (here and elsewhere) only serves the purpose of fixing -the name of the type parameter @{typ 'av} which would otherwise be renamed to -@{typ 'a}. *} - -locale Abs_Int = Gamma where \=\ for \ :: "'av::SL_top \ val set" -begin - -fun step' :: "'av st option \ 'av st option acom \ 'av st option acom" where -"step' S (SKIP {P}) = (SKIP {S})" | -"step' S (x ::= e {P}) = - x ::= e {case S of None \ None | Some S \ Some(update S x (aval' e S))}" | -"step' S (c1;; c2) = step' S c1;; step' (post c1) c2" | -"step' S (IF b THEN c1 ELSE c2 {P}) = - (let c1' = step' S c1; c2' = step' S c2 - in IF b THEN c1' ELSE c2' {post c1 \ post c2})" | -"step' S ({Inv} WHILE b DO c {P}) = - {S \ post c} WHILE b DO step' Inv c {Inv}" - -definition AI :: "com \ 'av st option acom option" where -"AI = lpfp\<^sub>c (step' \)" - - -lemma strip_step'[simp]: "strip(step' S c) = strip c" -by(induct c arbitrary: S) (simp_all add: Let_def) - - -text{* Soundness: *} - -lemma in_gamma_update: - "\ s : \\<^sub>f S; i : \ a \ \ s(x := i) : \\<^sub>f(update S x a)" -by(simp add: \_st_def lookup_update) - -text{* The soundness proofs are textually identical to the ones for the step -function operating on states as functions. *} - -lemma step_preserves_le: - "\ S \ \\<^sub>o S'; c \ \\<^sub>c c' \ \ step S c \ \\<^sub>c (step' S' c')" -proof(induction c arbitrary: c' S S') - case SKIP thus ?case by(auto simp:SKIP_le map_acom_SKIP) -next - case Assign thus ?case - by (fastforce simp: Assign_le map_acom_Assign intro: aval'_sound in_gamma_update - split: option.splits del:subsetD) -next - case Seq thus ?case apply (auto simp: Seq_le map_acom_Seq) - by (metis le_post post_map_acom) -next - case (If b c1 c2 P) - then obtain c1' c2' P' where - "c' = IF b THEN c1' ELSE c2' {P'}" - "P \ \\<^sub>o P'" "c1 \ \\<^sub>c c1'" "c2 \ \\<^sub>c c2'" - by (fastforce simp: If_le map_acom_If) - moreover have "post c1 \ \\<^sub>o(post c1' \ post c2')" - by (metis (no_types) `c1 \ \\<^sub>c c1'` join_ge1 le_post mono_gamma_o order_trans post_map_acom) - moreover have "post c2 \ \\<^sub>o(post c1' \ post c2')" - by (metis (no_types) `c2 \ \\<^sub>c c2'` join_ge2 le_post mono_gamma_o order_trans post_map_acom) - ultimately show ?case using `S \ \\<^sub>o S'` by (simp add: If.IH subset_iff) -next - case (While I b c1 P) - then obtain c1' I' P' where - "c' = {I'} WHILE b DO c1' {P'}" - "I \ \\<^sub>o I'" "P \ \\<^sub>o P'" "c1 \ \\<^sub>c c1'" - by (fastforce simp: map_acom_While While_le) - moreover have "S \ post c1 \ \\<^sub>o (S' \ post c1')" - using `S \ \\<^sub>o S'` le_post[OF `c1 \ \\<^sub>c c1'`, simplified] - by (metis (no_types) join_ge1 join_ge2 le_sup_iff mono_gamma_o order_trans) - ultimately show ?case by (simp add: While.IH subset_iff) -qed - -lemma AI_sound: "AI c = Some c' \ CS c \ \\<^sub>c c'" -proof(simp add: CS_def AI_def) - assume 1: "lpfp\<^sub>c (step' \) c = Some c'" - have 2: "step' \ c' \ c'" by(rule lpfpc_pfp[OF 1]) - have 3: "strip (\\<^sub>c (step' \ c')) = c" - by(simp add: strip_lpfpc[OF _ 1]) - have "lfp (step UNIV) c \ \\<^sub>c (step' \ c')" - proof(rule lfp_lowerbound[simplified,OF 3]) - show "step UNIV (\\<^sub>c (step' \ c')) \ \\<^sub>c (step' \ c')" - proof(rule step_preserves_le[OF _ _]) - show "UNIV \ \\<^sub>o \" by simp - show "\\<^sub>c (step' \ c') \ \\<^sub>c c'" by(rule mono_gamma_c[OF 2]) - qed - qed - from this 2 show "lfp (step UNIV) c \ \\<^sub>c c'" - by (blast intro: mono_gamma_c order_trans) -qed - -end - - -subsubsection "Monotonicity" - -locale Abs_Int_mono = Abs_Int + -assumes mono_plus': "a1 \ b1 \ a2 \ b2 \ plus' a1 a2 \ plus' b1 b2" -begin - -lemma mono_aval': "S \ S' \ aval' e S \ aval' e S'" -by(induction e) (auto simp: le_st_def lookup_def mono_plus') - -lemma mono_update: "a \ a' \ S \ S' \ update S x a \ update S' x a'" -by(auto simp add: le_st_def lookup_def update_def) - -lemma mono_step': "S \ S' \ c \ c' \ step' S c \ step' S' c'" -apply(induction c c' arbitrary: S S' rule: le_acom.induct) -apply (auto simp: Let_def mono_update mono_aval' mono_post le_join_disj - split: option.split) -done - -end - - -subsubsection "Ascending Chain Condition" - -abbreviation "strict r == r \ -(r^-1)" -abbreviation "acc r == wf((strict r)^-1)" - -lemma strict_inv_image: "strict(inv_image r f) = inv_image (strict r) f" -by(auto simp: inv_image_def) - -lemma acc_inv_image: - "acc r \ acc (inv_image r f)" -by (metis converse_inv_image strict_inv_image wf_inv_image) - -text{* ACC for option type: *} - -lemma acc_option: assumes "acc {(x,y::'a::preord). x \ y}" -shows "acc {(x,y::'a::preord option). x \ y}" -proof(auto simp: wf_eq_minimal) - fix xo :: "'a option" and Qo assume "xo : Qo" - let ?Q = "{x. Some x \ Qo}" - show "\yo\Qo. \zo. yo \ zo \ ~ zo \ yo \ zo \ Qo" (is "\zo\Qo. ?P zo") - proof cases - assume "?Q = {}" - hence "?P None" by auto - moreover have "None \ Qo" using `?Q = {}` `xo : Qo` - by auto (metis not_Some_eq) - ultimately show ?thesis by blast - next - assume "?Q \ {}" - with assms show ?thesis - apply(auto simp: wf_eq_minimal) - apply(erule_tac x="?Q" in allE) - apply auto - apply(rule_tac x = "Some z" in bexI) - by auto - qed -qed - -text{* ACC for abstract states, via measure functions. *} - -lemma measure_st: assumes "(strict{(x,y::'a::SL_top). x \ y})^-1 <= measure m" -and "\x y::'a::SL_top. x \ y \ y \ x \ m x = m y" -shows "(strict{(S,S'::'a::SL_top st). S \ S'})^-1 \ - measure(%fd. \x| x\set(dom fd) \ ~ \ \ fun fd x. m(fun fd x)+1)" -proof- - { fix S S' :: "'a st" assume "S \ S'" "~ S' \ S" - let ?X = "set(dom S)" let ?Y = "set(dom S')" - let ?f = "fun S" let ?g = "fun S'" - let ?X' = "{x:?X. ~ \ \ ?f x}" let ?Y' = "{y:?Y. ~ \ \ ?g y}" - from `S \ S'` have "ALL y:?Y'\?X. ?f y \ ?g y" - by(auto simp: le_st_def lookup_def) - hence 1: "ALL y:?Y'\?X. m(?g y)+1 \ m(?f y)+1" - using assms(1,2) by(fastforce) - from `~ S' \ S` obtain u where u: "u : ?X" "~ lookup S' u \ ?f u" - by(auto simp: le_st_def) - hence "u : ?X'" by simp (metis preord_class.le_trans top) - have "?Y'-?X = {}" using `S \ S'` by(fastforce simp: le_st_def lookup_def) - have "?Y'\?X <= ?X'" apply auto - apply (metis `S \ S'` le_st_def lookup_def preord_class.le_trans) - done - have "(\y\?Y'. m(?g y)+1) = (\y\(?Y'-?X) \ (?Y'\?X). m(?g y)+1)" - by (metis Un_Diff_Int) - also have "\ = (\y\?Y'\?X. m(?g y)+1)" - using `?Y'-?X = {}` by (metis Un_empty_left) - also have "\ < (\x\?X'. m(?f x)+1)" - proof cases - assume "u \ ?Y'" - hence "m(?g u) < m(?f u)" using assms(1) `S \ S'` u - by (fastforce simp: le_st_def lookup_def) - have "(\y\?Y'\?X. m(?g y)+1) < (\y\?Y'\?X. m(?f y)+1)" - using `u:?X` `u:?Y'` `m(?g u) < m(?f u)` - by(fastforce intro!: sum_strict_mono_ex1[OF _ 1]) - also have "\ \ (\y\?X'. m(?f y)+1)" - by(simp add: sum_mono3[OF _ `?Y'\?X <= ?X'`]) - finally show ?thesis . - next - assume "u \ ?Y'" - with `?Y'\?X <= ?X'` have "?Y'\?X - {u} <= ?X' - {u}" by blast - have "(\y\?Y'\?X. m(?g y)+1) = (\y\?Y'\?X - {u}. m(?g y)+1)" - proof- - have "?Y'\?X = ?Y'\?X - {u}" using `u \ ?Y'` by auto - thus ?thesis by metis - qed - also have "\ < (\y\?Y'\?X-{u}. m(?g y)+1) + (\y\{u}. m(?f y)+1)" by simp - also have "(\y\?Y'\?X-{u}. m(?g y)+1) \ (\y\?Y'\?X-{u}. m(?f y)+1)" - using 1 by(blast intro: sum_mono) - also have "\ \ (\y\?X'-{u}. m(?f y)+1)" - by(simp add: sum_mono3[OF _ `?Y'\?X-{u} <= ?X'-{u}`]) - also have "\ + (\y\{u}. m(?f y)+1)= (\y\(?X'-{u}) \ {u}. m(?f y)+1)" - using `u:?X'` by(subst sum.union_disjoint[symmetric]) auto - also have "\ = (\x\?X'. m(?f x)+1)" - using `u : ?X'` by(simp add:insert_absorb) - finally show ?thesis by (blast intro: add_right_mono) - qed - finally have "(\y\?Y'. m(?g y)+1) < (\x\?X'. m(?f x)+1)" . - } thus ?thesis by(auto simp add: measure_def inv_image_def) -qed - -text{* ACC for acom. First the ordering on acom is related to an ordering on -lists of annotations. *} - -(* FIXME mv and add [simp] *) -lemma listrel_Cons_iff: - "(x#xs, y#ys) : listrel r \ (x,y) \ r \ (xs,ys) \ listrel r" -by (blast intro:listrel.Cons) - -lemma listrel_app: "(xs1,ys1) : listrel r \ (xs2,ys2) : listrel r - \ (xs1@xs2, ys1@ys2) : listrel r" -by(auto simp add: listrel_iff_zip) - -lemma listrel_app_same_size: "size xs1 = size ys1 \ size xs2 = size ys2 \ - (xs1@xs2, ys1@ys2) : listrel r \ - (xs1,ys1) : listrel r \ (xs2,ys2) : listrel r" -by(auto simp add: listrel_iff_zip) - -lemma listrel_converse: "listrel(r^-1) = (listrel r)^-1" -proof- - { fix xs ys - have "(xs,ys) : listrel(r^-1) \ (ys,xs) : listrel r" - apply(induct xs arbitrary: ys) - apply (fastforce simp: listrel.Nil) - apply (fastforce simp: listrel_Cons_iff) - done - } thus ?thesis by auto -qed - -(* It would be nice to get rid of refl & trans and build them into the proof *) -lemma acc_listrel: fixes r :: "('a*'a)set" assumes "refl r" and "trans r" -and "acc r" shows "acc (listrel r - {([],[])})" -proof- - have refl: "!!x. (x,x) : r" using `refl r` unfolding refl_on_def by blast - have trans: "!!x y z. (x,y) : r \ (y,z) : r \ (x,z) : r" - using `trans r` unfolding trans_def by blast - from assms(3) obtain mx :: "'a set \ 'a" where - mx: "!!S x. x:S \ mx S : S \ (\y. (mx S,y) : strict r \ y \ S)" - by(simp add: wf_eq_minimal) metis - let ?R = "listrel r - {([], [])}" - { fix Q and xs :: "'a list" - have "xs \ Q \ \ys. ys\Q \ (\zs. (ys, zs) \ strict ?R \ zs \ Q)" - (is "_ \ \ys. ?P Q ys") - proof(induction xs arbitrary: Q rule: length_induct) - case (1 xs) - { have "!!ys Q. size ys < size xs \ ys : Q \ EX ms. ?P Q ms" - using "1.IH" by blast - } note IH = this - show ?case - proof(cases xs) - case Nil with `xs : Q` have "?P Q []" by auto - thus ?thesis by blast - next - case (Cons x ys) - let ?Q1 = "{a. \bs. size bs = size ys \ a#bs : Q}" - have "x : ?Q1" using `xs : Q` Cons by auto - from mx[OF this] obtain m1 where - 1: "m1 \ ?Q1 \ (\y. (m1,y) \ strict r \ y \ ?Q1)" by blast - then obtain ms1 where "size ms1 = size ys" "m1#ms1 : Q" by blast+ - hence "size ms1 < size xs" using Cons by auto - let ?Q2 = "{bs. \m1'. (m1',m1):r \ (m1,m1'):r \ m1'#bs : Q \ size bs = size ms1}" - have "ms1 : ?Q2" using `m1#ms1 : Q` by(blast intro: refl) - from IH[OF `size ms1 < size xs` this] - obtain ms where 2: "?P ?Q2 ms" by auto - then obtain m1' where m1': "(m1',m1) : r \ (m1,m1') : r \ m1'#ms : Q" - by blast - hence "\ab. (m1'#ms,ab) : strict ?R \ ab \ Q" using 1 2 - apply (auto simp: listrel_Cons_iff) - apply (metis `length ms1 = length ys` listrel_eq_len trans) - by (metis `length ms1 = length ys` listrel_eq_len trans) - with m1' show ?thesis by blast - qed - qed - } - thus ?thesis unfolding wf_eq_minimal by (metis converse_iff) -qed - - -lemma le_iff_le_annos: "c1 \ c2 \ - (annos c1, annos c2) : listrel{(x,y). x \ y} \ strip c1 = strip c2" -apply(induct c1 c2 rule: le_acom.induct) -apply (auto simp: listrel.Nil listrel_Cons_iff listrel_app size_annos_same2) -apply (metis listrel_app_same_size size_annos_same)+ -done - -lemma le_acom_subset_same_annos: - "(strict{(c,c'::'a::preord acom). c \ c'})^-1 \ - (strict(inv_image (listrel{(a,a'::'a). a \ a'} - {([],[])}) annos))^-1" -by(auto simp: le_iff_le_annos) - -lemma acc_acom: "acc {(a,a'::'a::preord). a \ a'} \ - acc {(c,c'::'a acom). c \ c'}" -apply(rule wf_subset[OF _ le_acom_subset_same_annos]) -apply(rule acc_inv_image[OF acc_listrel]) -apply(auto simp: refl_on_def trans_def intro: le_trans) -done - -text{* Termination of the fixed-point finders, assuming monotone functions: *} - -lemma pfp_termination: -fixes x0 :: "'a::preord" -assumes mono: "\x y. x \ y \ f x \ f y" and "acc {(x::'a,y). x \ y}" -and "x0 \ f x0" shows "EX x. pfp f x0 = Some x" -proof(simp add: pfp_def, rule wf_while_option_Some[where P = "%x. x \ f x"]) - show "wf {(x, s). (s \ f s \ \ f s \ s) \ x = f s}" - by(rule wf_subset[OF assms(2)]) auto -next - show "x0 \ f x0" by(rule assms) -next - fix x assume "x \ f x" thus "f x \ f(f x)" by(rule mono) -qed - -lemma lpfpc_termination: - fixes f :: "(('a::SL_top)option acom \ 'a option acom)" - assumes "acc {(x::'a,y). x \ y}" and "\x y. x \ y \ f x \ f y" - and "\c. strip(f c) = strip c" - shows "\c'. lpfp\<^sub>c f c = Some c'" -unfolding lpfp\<^sub>c_def -apply(rule pfp_termination) - apply(erule assms(2)) - apply(rule acc_acom[OF acc_option[OF assms(1)]]) -apply(simp add: bot_acom assms(3)) -done - -context Abs_Int_mono -begin - -lemma AI_Some_measure: -assumes "(strict{(x,y::'a). x \ y})^-1 <= measure m" -and "\x y::'a. x \ y \ y \ x \ m x = m y" -shows "\c'. AI c = Some c'" -unfolding AI_def -apply(rule lpfpc_termination) -apply(rule wf_subset[OF wf_measure measure_st[OF assms]]) -apply(erule mono_step'[OF le_refl]) -apply(rule strip_step') -done - -end - -end diff -r e44f5c123f26 -r 49b78f1f9e01 src/HOL/IMP/Abs_Int_ITP/Abs_Int1_const_ITP.thy --- a/src/HOL/IMP/Abs_Int_ITP/Abs_Int1_const_ITP.thy Tue Nov 22 18:36:59 2016 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,139 +0,0 @@ -(* Author: Tobias Nipkow *) - -theory Abs_Int1_const_ITP -imports Abs_Int1_ITP "../Abs_Int_Tests" -begin - -subsection "Constant Propagation" - -datatype const = Const val | Any - -fun \_const where -"\_const (Const n) = {n}" | -"\_const (Any) = UNIV" - -fun plus_const where -"plus_const (Const m) (Const n) = Const(m+n)" | -"plus_const _ _ = Any" - -lemma plus_const_cases: "plus_const a1 a2 = - (case (a1,a2) of (Const m, Const n) \ Const(m+n) | _ \ Any)" -by(auto split: prod.split const.split) - -instantiation const :: SL_top -begin - -fun le_const where -"_ \ Any = True" | -"Const n \ Const m = (n=m)" | -"Any \ Const _ = False" - -fun join_const where -"Const m \ Const n = (if n=m then Const m else Any)" | -"_ \ _ = Any" - -definition "\ = 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_const_def) -qed - -end - - -global_interpretation Val_abs -where \ = \_const and num' = Const and plus' = plus_const -proof - case goal1 thus ?case - by(cases a, cases b, simp, simp, cases b, simp, simp) -next - case goal2 show ?case by(simp add: Top_const_def) -next - case goal3 show ?case by simp -next - case goal4 thus ?case - by(auto simp: plus_const_cases split: const.split) -qed - -global_interpretation Abs_Int -where \ = \_const and num' = Const and plus' = plus_const -defines AI_const = AI and step_const = step' and aval'_const = aval' -.. - - -subsubsection "Tests" - -value "show_acom (((step_const \)^^0) (\\<^sub>c test1_const))" -value "show_acom (((step_const \)^^1) (\\<^sub>c test1_const))" -value "show_acom (((step_const \)^^2) (\\<^sub>c test1_const))" -value "show_acom (((step_const \)^^3) (\\<^sub>c test1_const))" -value "show_acom_opt (AI_const test1_const)" - -value "show_acom_opt (AI_const test2_const)" -value "show_acom_opt (AI_const test3_const)" - -value "show_acom (((step_const \)^^0) (\\<^sub>c test4_const))" -value "show_acom (((step_const \)^^1) (\\<^sub>c test4_const))" -value "show_acom (((step_const \)^^2) (\\<^sub>c test4_const))" -value "show_acom (((step_const \)^^3) (\\<^sub>c test4_const))" -value "show_acom_opt (AI_const test4_const)" - -value "show_acom (((step_const \)^^0) (\\<^sub>c test5_const))" -value "show_acom (((step_const \)^^1) (\\<^sub>c test5_const))" -value "show_acom (((step_const \)^^2) (\\<^sub>c test5_const))" -value "show_acom (((step_const \)^^3) (\\<^sub>c test5_const))" -value "show_acom (((step_const \)^^4) (\\<^sub>c test5_const))" -value "show_acom (((step_const \)^^5) (\\<^sub>c test5_const))" -value "show_acom_opt (AI_const test5_const)" - -value "show_acom (((step_const \)^^0) (\\<^sub>c test6_const))" -value "show_acom (((step_const \)^^1) (\\<^sub>c test6_const))" -value "show_acom (((step_const \)^^2) (\\<^sub>c test6_const))" -value "show_acom (((step_const \)^^3) (\\<^sub>c test6_const))" -value "show_acom (((step_const \)^^4) (\\<^sub>c test6_const))" -value "show_acom (((step_const \)^^5) (\\<^sub>c test6_const))" -value "show_acom (((step_const \)^^6) (\\<^sub>c test6_const))" -value "show_acom (((step_const \)^^7) (\\<^sub>c test6_const))" -value "show_acom (((step_const \)^^8) (\\<^sub>c test6_const))" -value "show_acom (((step_const \)^^9) (\\<^sub>c test6_const))" -value "show_acom (((step_const \)^^10) (\\<^sub>c test6_const))" -value "show_acom (((step_const \)^^11) (\\<^sub>c test6_const))" -value "show_acom_opt (AI_const test6_const)" - - -text{* Monotonicity: *} - -global_interpretation Abs_Int_mono -where \ = \_const and num' = Const and plus' = plus_const -proof - case goal1 thus ?case - by(auto simp: plus_const_cases split: const.split) -qed - -text{* Termination: *} - -definition "m_const x = (case x of Const _ \ 1 | Any \ 0)" - -lemma measure_const: - "(strict{(x::const,y). x \ y})^-1 \ measure m_const" -by(auto simp: m_const_def split: const.splits) - -lemma measure_const_eq: - "\ x y::const. x \ y \ y \ x \ m_const x = m_const y" -by(auto simp: m_const_def split: const.splits) - -lemma "EX c'. AI_const c = Some c'" -by(rule AI_Some_measure[OF measure_const measure_const_eq]) - -end diff -r e44f5c123f26 -r 49b78f1f9e01 src/HOL/IMP/Abs_Int_ITP/Abs_Int1_parity_ITP.thy --- a/src/HOL/IMP/Abs_Int_ITP/Abs_Int1_parity_ITP.thy Tue Nov 22 18:36:59 2016 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,168 +0,0 @@ -(* Author: Tobias Nipkow *) - -theory Abs_Int1_parity_ITP -imports Abs_Int1_ITP -begin - -subsection "Parity Analysis" - -datatype parity = Even | Odd | Either - -text{* Instantiation of class @{class preord} with type @{typ parity}: *} - -instantiation parity :: preord -begin - -text{* First the definition of the interface function @{text"\"}. Note that -the header of the definition must refer to the ascii name @{const le} of the -constants as @{text le_parity} and the definition is named @{text -le_parity_def}. Inside the definition the symbolic names can be used. *} - -definition le_parity where -"x \ y = (y = Either \ x=y)" - -text{* Now the instance proof, i.e.\ the proof that the definition fulfills -the axioms (assumptions) of the class. The initial proof-step generates the -necessary proof obligations. *} - -instance -proof - fix x::parity show "x \ x" by(auto simp: le_parity_def) -next - fix x y z :: parity assume "x \ y" "y \ z" thus "x \ z" - by(auto simp: le_parity_def) -qed - -end - -text{* Instantiation of class @{class SL_top} with type @{typ parity}: *} - -instantiation parity :: SL_top -begin - - -definition join_parity where -"x \ y = (if x \ y then y else if y \ x then x else Either)" - -definition Top_parity where -"\ = Either" - -text{* Now the instance proof. This time we take a lazy shortcut: we do not -write out the proof obligations but use the @{text goali} primitive to refer -to the assumptions of subgoal i and @{text "case?"} to refer to the -conclusion of subgoal i. The class axioms are presented in the same order as -in the class definition. *} - -instance -proof - case goal1 (*join1*) show ?case by(auto simp: le_parity_def join_parity_def) -next - case goal2 (*join2*) show ?case by(auto simp: le_parity_def join_parity_def) -next - case goal3 (*join least*) thus ?case by(auto simp: le_parity_def join_parity_def) -next - case goal4 (*Top*) show ?case by(auto simp: le_parity_def Top_parity_def) -qed - -end - - -text{* Now we define the functions used for instantiating the abstract -interpretation locales. Note that the Isabelle terminology is -\emph{interpretation}, not \emph{instantiation} of locales, but we use -instantiation to avoid confusion with abstract interpretation. *} - -fun \_parity :: "parity \ val set" where -"\_parity Even = {i. i mod 2 = 0}" | -"\_parity Odd = {i. i mod 2 = 1}" | -"\_parity Either = UNIV" - -fun num_parity :: "val \ parity" where -"num_parity i = (if i mod 2 = 0 then Even else Odd)" - -fun plus_parity :: "parity \ parity \ parity" where -"plus_parity Even Even = Even" | -"plus_parity Odd Odd = Even" | -"plus_parity Even Odd = Odd" | -"plus_parity Odd Even = Odd" | -"plus_parity Either y = Either" | -"plus_parity x Either = Either" - -text{* First we instantiate the abstract value interface and prove that the -functions on type @{typ parity} have all the necessary properties: *} - -interpretation Val_abs -where \ = \_parity and num' = num_parity and plus' = plus_parity -proof txt{* of the locale axioms *} - fix a b :: parity - assume "a \ b" thus "\_parity a \ \_parity b" - by(auto simp: le_parity_def) -next txt{* The rest in the lazy, implicit way *} - case goal2 show ?case by(auto simp: Top_parity_def) -next - case goal3 show ?case by auto -next - txt{* Warning: this subproof refers to the names @{text a1} and @{text a2} - from the statement of the axiom. *} - case goal4 thus ?case - proof(cases a1 a2 rule: parity.exhaust[case_product parity.exhaust]) - qed (auto simp add:mod_add_eq) -qed - -text{* Instantiating the abstract interpretation locale requires no more -proofs (they happened in the instatiation above) but delivers the -instantiated abstract interpreter which we call AI: *} - -global_interpretation Abs_Int -where \ = \_parity and num' = num_parity and plus' = plus_parity -defines aval_parity = aval' and step_parity = step' and AI_parity = AI -.. - - -subsubsection "Tests" - -definition "test1_parity = - ''x'' ::= N 1;; - WHILE Less (V ''x'') (N 100) DO ''x'' ::= Plus (V ''x'') (N 2)" - -value "show_acom_opt (AI_parity test1_parity)" - -definition "test2_parity = - ''x'' ::= N 1;; - WHILE Less (V ''x'') (N 100) DO ''x'' ::= Plus (V ''x'') (N 3)" - -value "show_acom ((step_parity \ ^^1) (anno None test2_parity))" -value "show_acom ((step_parity \ ^^2) (anno None test2_parity))" -value "show_acom ((step_parity \ ^^3) (anno None test2_parity))" -value "show_acom ((step_parity \ ^^4) (anno None test2_parity))" -value "show_acom ((step_parity \ ^^5) (anno None test2_parity))" -value "show_acom_opt (AI_parity test2_parity)" - - -subsubsection "Termination" - -global_interpretation Abs_Int_mono -where \ = \_parity and num' = num_parity and plus' = plus_parity -proof - case goal1 thus ?case - proof(cases a1 a2 b1 b2 - rule: parity.exhaust[case_product parity.exhaust[case_product parity.exhaust[case_product parity.exhaust]]]) (* FIXME - UGLY! *) - qed (auto simp add:le_parity_def) -qed - - -definition m_parity :: "parity \ nat" where -"m_parity x = (if x=Either then 0 else 1)" - -lemma measure_parity: - "(strict{(x::parity,y). x \ y})^-1 \ measure m_parity" -by(auto simp add: m_parity_def le_parity_def) - -lemma measure_parity_eq: - "\x y::parity. x \ y \ y \ x \ m_parity x = m_parity y" -by(auto simp add: m_parity_def le_parity_def) - -lemma AI_parity_Some: "\c'. AI_parity c = Some c'" -by(rule AI_Some_measure[OF measure_parity measure_parity_eq]) - -end diff -r e44f5c123f26 -r 49b78f1f9e01 src/HOL/IMP/Abs_Int_ITP/Abs_Int2_ITP.thy --- a/src/HOL/IMP/Abs_Int_ITP/Abs_Int2_ITP.thy Tue Nov 22 18:36:59 2016 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,340 +0,0 @@ -(* Author: Tobias Nipkow *) - -theory Abs_Int2_ITP -imports Abs_Int1_ITP "../Vars" -begin - -instantiation prod :: (preord,preord) preord -begin - -definition "le_prod p1 p2 = (fst p1 \ fst p2 \ snd p1 \ snd p2)" - -instance -proof - case goal1 show ?case by(simp add: le_prod_def) -next - case goal2 thus ?case unfolding le_prod_def by(metis le_trans) -qed - -end - - -subsection "Backward Analysis of Expressions" - -hide_const bot - -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]: "\ \ x" -begin - -lemma mono_meet: "x \ x' \ y \ y' \ x \ y \ x' \ y'" -by (metis meet_greatest meet_le1 meet_le2 le_trans) - -end - -locale Val_abs1_gamma = - Gamma where \ = \ for \ :: "'av::L_top_bot \ val set" + -assumes inter_gamma_subset_gamma_meet: - "\ a1 \ \ a2 \ \(a1 \ a2)" -and gamma_Bot[simp]: "\ \ = {}" -begin - -lemma in_gamma_meet: "x : \ a1 \ x : \ a2 \ x : \(a1 \ a2)" -by (metis IntI inter_gamma_subset_gamma_meet set_mp) - -lemma gamma_meet[simp]: "\(a1 \ a2) = \ a1 \ \ a2" -by (metis equalityI inter_gamma_subset_gamma_meet le_inf_iff mono_gamma meet_le1 meet_le2) - -end - - -locale Val_abs1 = - Val_abs1_gamma where \ = \ - for \ :: "'av::L_top_bot \ val set" + -fixes test_num' :: "val \ 'av \ bool" -and filter_plus' :: "'av \ 'av \ 'av \ 'av * 'av" -and filter_less' :: "bool \ 'av \ 'av \ 'av * 'av" -assumes test_num': "test_num' n a = (n : \ a)" -and filter_plus': "filter_plus' a a1 a2 = (b1,b2) \ - n1 : \ a1 \ n2 : \ a2 \ n1+n2 : \ a \ n1 : \ b1 \ n2 : \ b2" -and filter_less': "filter_less' (n1 - n1 : \ a1 \ n2 : \ a2 \ n1 : \ b1 \ n2 : \ b2" - - -locale Abs_Int1 = - Val_abs1 where \ = \ for \ :: "'av::L_top_bot \ val set" -begin - -lemma in_gamma_join_UpI: "s : \\<^sub>o S1 \ s : \\<^sub>o S2 \ s : \\<^sub>o(S1 \ S2)" -by (metis (no_types) join_ge1 join_ge2 mono_gamma_o set_rev_mp) - -fun aval'' :: "aexp \ 'av st option \ 'av" where -"aval'' e None = \" | -"aval'' e (Some sa) = aval' e sa" - -lemma aval''_sound: "s : \\<^sub>o S \ aval a s : \(aval'' a S)" -by(cases S)(simp add: aval'_sound)+ - -subsubsection "Backward analysis" - -fun afilter :: "aexp \ 'av \ 'av st option \ 'av st option" where -"afilter (N n) a S = (if test_num' n a then S else None)" | -"afilter (V x) a S = (case S of None \ None | Some S \ - let a' = lookup S x \ a in - if a' \ \ then None else Some(update S x a'))" | -"afilter (Plus e1 e2) a S = - (let (a1,a2) = filter_plus' a (aval'' e1 S) (aval'' e2 S) - in afilter e1 a1 (afilter e2 a2 S))" - -text{* The test for @{const bot} in the @{const V}-case is important: @{const -bot} indicates that a variable has no possible values, i.e.\ that the current -program point is unreachable. But then the abstract state should collapse to -@{const None}. Put differently, we maintain the invariant that in an abstract -state of the form @{term"Some s"}, all variables are mapped to non-@{const -bot} values. Otherwise the (pointwise) join of two abstract states, one of -which contains @{const bot} values, may produce too large a result, thus -making the analysis less precise. *} - - -fun bfilter :: "bexp \ bool \ 'av st option \ 'av st option" where -"bfilter (Bc v) res S = (if v=res then S else None)" | -"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) = filter_less' res (aval'' e1 S) (aval'' e2 S) - in afilter e1 res1 (afilter e2 res2 S))" - -lemma afilter_sound: "s : \\<^sub>o S \ aval e s : \ a \ s : \\<^sub>o (afilter e a S)" -proof(induction e arbitrary: a S) - case N thus ?case by simp (metis test_num') -next - case (V x) - obtain S' where "S = Some S'" and "s : \\<^sub>f S'" using `s : \\<^sub>o S` - by(auto simp: in_gamma_option_iff) - moreover hence "s x : \ (lookup S' x)" by(simp add: \_st_def) - moreover have "s x : \ a" using V by simp - ultimately show ?case using V(1) - by(simp add: lookup_update Let_def \_st_def) - (metis mono_gamma emptyE in_gamma_meet gamma_Bot subset_empty) -next - case (Plus e1 e2) thus ?case - using filter_plus'[OF _ aval''_sound[OF Plus(3)] aval''_sound[OF Plus(3)]] - by (auto split: prod.split) -qed - -lemma bfilter_sound: "s : \\<^sub>o S \ bv = bval b s \ s : \\<^sub>o(bfilter b bv S)" -proof(induction b arbitrary: S bv) - case Bc thus ?case by simp -next - case (Not b) thus ?case by simp -next - case (And b1 b2) thus ?case - apply hypsubst_thin - apply (fastforce simp: in_gamma_join_UpI) - done -next - case (Less e1 e2) thus ?case - apply hypsubst_thin - apply (auto split: prod.split) - apply (metis afilter_sound filter_less' aval''_sound Less) - done -qed - - -fun step' :: "'av st option \ 'av st option acom \ 'av st option acom" - where -"step' S (SKIP {P}) = (SKIP {S})" | -"step' S (x ::= e {P}) = - x ::= e {case S of None \ None | Some S \ Some(update S x (aval' e S))}" | -"step' S (c1;; c2) = step' S c1;; step' (post c1) c2" | -"step' S (IF b THEN c1 ELSE c2 {P}) = - (let c1' = step' (bfilter b True S) c1; c2' = step' (bfilter b False S) c2 - in IF b THEN c1' ELSE c2' {post c1 \ post c2})" | -"step' S ({Inv} WHILE b DO c {P}) = - {S \ post c} - WHILE b DO step' (bfilter b True Inv) c - {bfilter b False Inv}" - -definition AI :: "com \ 'av st option acom option" where -"AI = lpfp\<^sub>c (step' \)" - -lemma strip_step'[simp]: "strip(step' S c) = strip c" -by(induct c arbitrary: S) (simp_all add: Let_def) - - -subsubsection "Soundness" - -lemma in_gamma_update: - "\ s : \\<^sub>f S; i : \ a \ \ s(x := i) : \\<^sub>f(update S x a)" -by(simp add: \_st_def lookup_update) - -lemma step_preserves_le: - "\ S \ \\<^sub>o S'; cs \ \\<^sub>c ca \ \ step S cs \ \\<^sub>c (step' S' ca)" -proof(induction cs arbitrary: ca S S') - case SKIP thus ?case by(auto simp:SKIP_le map_acom_SKIP) -next - case Assign thus ?case - by (fastforce simp: Assign_le map_acom_Assign intro: aval'_sound in_gamma_update - split: option.splits del:subsetD) -next - case Seq thus ?case apply (auto simp: Seq_le map_acom_Seq) - by (metis le_post post_map_acom) -next - case (If b cs1 cs2 P) - then obtain ca1 ca2 Pa where - "ca= IF b THEN ca1 ELSE ca2 {Pa}" - "P \ \\<^sub>o Pa" "cs1 \ \\<^sub>c ca1" "cs2 \ \\<^sub>c ca2" - by (fastforce simp: If_le map_acom_If) - moreover have "post cs1 \ \\<^sub>o(post ca1 \ post ca2)" - by (metis (no_types) `cs1 \ \\<^sub>c ca1` join_ge1 le_post mono_gamma_o order_trans post_map_acom) - moreover have "post cs2 \ \\<^sub>o(post ca1 \ post ca2)" - by (metis (no_types) `cs2 \ \\<^sub>c ca2` join_ge2 le_post mono_gamma_o order_trans post_map_acom) - ultimately show ?case using `S \ \\<^sub>o S'` - by (simp add: If.IH subset_iff bfilter_sound) -next - case (While I b cs1 P) - then obtain ca1 Ia Pa where - "ca = {Ia} WHILE b DO ca1 {Pa}" - "I \ \\<^sub>o Ia" "P \ \\<^sub>o Pa" "cs1 \ \\<^sub>c ca1" - by (fastforce simp: map_acom_While While_le) - moreover have "S \ post cs1 \ \\<^sub>o (S' \ post ca1)" - using `S \ \\<^sub>o S'` le_post[OF `cs1 \ \\<^sub>c ca1`, simplified] - by (metis (no_types) join_ge1 join_ge2 le_sup_iff mono_gamma_o order_trans) - ultimately show ?case by (simp add: While.IH subset_iff bfilter_sound) -qed - -lemma AI_sound: "AI c = Some c' \ CS c \ \\<^sub>c c'" -proof(simp add: CS_def AI_def) - assume 1: "lpfp\<^sub>c (step' \) c = Some c'" - have 2: "step' \ c' \ c'" by(rule lpfpc_pfp[OF 1]) - have 3: "strip (\\<^sub>c (step' \ c')) = c" - by(simp add: strip_lpfpc[OF _ 1]) - have "lfp (step UNIV) c \ \\<^sub>c (step' \ c')" - proof(rule lfp_lowerbound[simplified,OF 3]) - show "step UNIV (\\<^sub>c (step' \ c')) \ \\<^sub>c (step' \ c')" - proof(rule step_preserves_le[OF _ _]) - show "UNIV \ \\<^sub>o \" by simp - show "\\<^sub>c (step' \ c') \ \\<^sub>c c'" by(rule mono_gamma_c[OF 2]) - qed - qed - from this 2 show "lfp (step UNIV) c \ \\<^sub>c c'" - by (blast intro: mono_gamma_c order_trans) -qed - - -subsubsection "Commands over a set of variables" - -text{* Key invariant: the domains of all abstract states are subsets of the -set of variables of the program. *} - -definition "domo S = (case S of None \ {} | Some S' \ set(dom S'))" - -definition Com :: "vname set \ 'a st option acom set" where -"Com X = {c. \S \ set(annos c). domo S \ X}" - -lemma domo_Top[simp]: "domo \ = {}" -by(simp add: domo_def Top_st_def Top_option_def) - -lemma bot_acom_Com[simp]: "\\<^sub>c c \ Com X" -by(simp add: bot_acom_def Com_def domo_def set_annos_anno) - -lemma post_in_annos: "post c : set(annos c)" -by(induction c) simp_all - -lemma domo_join: "domo (S \ T) \ domo S \ domo T" -by(auto simp: domo_def join_st_def split: option.split) - -lemma domo_afilter: "vars a \ X \ domo S \ X \ domo(afilter a i S) \ X" -apply(induction a arbitrary: i S) -apply(simp add: domo_def) -apply(simp add: domo_def Let_def update_def lookup_def split: option.splits) -apply blast -apply(simp split: prod.split) -done - -lemma domo_bfilter: "vars b \ X \ domo S \ X \ domo(bfilter b bv S) \ X" -apply(induction b arbitrary: bv S) -apply(simp add: domo_def) -apply(simp) -apply(simp) -apply rule -apply (metis le_sup_iff subset_trans[OF domo_join]) -apply(simp split: prod.split) -by (metis domo_afilter) - -lemma step'_Com: - "domo S \ X \ vars(strip c) \ X \ c : Com X \ step' S c : Com X" -apply(induction c arbitrary: S) -apply(simp add: Com_def) -apply(simp add: Com_def domo_def update_def split: option.splits) -apply(simp (no_asm_use) add: Com_def ball_Un) -apply (metis post_in_annos) -apply(simp (no_asm_use) add: Com_def ball_Un) -apply rule -apply (metis Un_assoc domo_join order_trans post_in_annos subset_Un_eq) -apply (metis domo_bfilter) -apply(simp (no_asm_use) add: Com_def) -apply rule -apply (metis domo_join le_sup_iff post_in_annos subset_trans) -apply rule -apply (metis domo_bfilter) -by (metis domo_bfilter) - -end - - -subsubsection "Monotonicity" - -locale Abs_Int1_mono = Abs_Int1 + -assumes mono_plus': "a1 \ b1 \ a2 \ b2 \ plus' a1 a2 \ plus' b1 b2" -and mono_filter_plus': "a1 \ b1 \ a2 \ b2 \ r \ r' \ - filter_plus' r a1 a2 \ filter_plus' r' b1 b2" -and mono_filter_less': "a1 \ b1 \ a2 \ b2 \ - filter_less' bv a1 a2 \ filter_less' bv b1 b2" -begin - -lemma mono_aval': "S \ S' \ aval' e S \ aval' e S'" -by(induction e) (auto simp: le_st_def lookup_def mono_plus') - -lemma mono_aval'': "S \ S' \ aval'' e S \ aval'' e S'" -apply(cases S) - apply simp -apply(cases S') - apply simp -by (simp add: mono_aval') - -lemma mono_afilter: "r \ r' \ S \ S' \ afilter e r S \ afilter e r' S'" -apply(induction e arbitrary: r r' S S') -apply(auto simp: test_num' Let_def split: option.splits prod.splits) -apply (metis mono_gamma subsetD) -apply(rename_tac list a b c d, drule_tac x = "list" in mono_lookup) -apply (metis mono_meet le_trans) -apply (metis mono_meet mono_lookup mono_update) -apply(metis mono_aval'' mono_filter_plus'[simplified le_prod_def] fst_conv snd_conv) -done - -lemma mono_bfilter: "S \ S' \ bfilter b r S \ bfilter b r S'" -apply(induction b arbitrary: r S S') -apply(auto simp: le_trans[OF _ join_ge1] le_trans[OF _ join_ge2] split: prod.splits) -apply(metis mono_aval'' mono_afilter mono_filter_less'[simplified le_prod_def] fst_conv snd_conv) -done - -lemma mono_step': "S \ S' \ c \ c' \ step' S c \ step' S' c'" -apply(induction c c' arbitrary: S S' rule: le_acom.induct) -apply (auto simp: mono_post mono_bfilter mono_update mono_aval' Let_def le_join_disj - split: option.split) -done - -lemma mono_step'2: "mono (step' S)" -by(simp add: mono_def mono_step'[OF le_refl]) - -end - -end diff -r e44f5c123f26 -r 49b78f1f9e01 src/HOL/IMP/Abs_Int_ITP/Abs_Int2_ivl_ITP.thy --- a/src/HOL/IMP/Abs_Int_ITP/Abs_Int2_ivl_ITP.thy Tue Nov 22 18:36:59 2016 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,285 +0,0 @@ -(* Author: Tobias Nipkow *) - -theory Abs_Int2_ivl_ITP -imports Abs_Int2_ITP "../Abs_Int_Tests" -begin - -subsection "Interval Analysis" - -datatype ivl = I "int option" "int option" - -definition "\_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)" - -abbreviation I_Some_Some :: "int \ int \ ivl" ("{_\_}") where -"{lo\hi} == I (Some lo) (Some hi)" -abbreviation I_Some_None :: "int \ ivl" ("{_\}") where -"{lo\} == I (Some lo) None" -abbreviation I_None_Some :: "int \ ivl" ("{\_}") where -"{\hi} == I None (Some hi)" -abbreviation I_None_None :: "ivl" ("{\}") where -"{\} == I None None" - -definition "num_ivl n = {n\n}" - -fun in_ivl :: "int \ ivl \ bool" where -"in_ivl k (I (Some l) (Some h)) \ l \ k \ k \ h" | -"in_ivl k (I (Some l) None) \ l \ k" | -"in_ivl k (I None (Some h)) \ k \ h" | -"in_ivl k (I None None) \ True" - -instantiation option :: (plus)plus -begin - -fun plus_option where -"Some x + Some y = Some(x+y)" | -"_ + _ = None" - -instance .. - -end - -definition empty where "empty = {1\0}" - -fun is_empty where -"is_empty {l\h} = (h (case h of Some h \ h False) | None \ False)" -by(auto split:option.split) - -lemma [simp]: "is_empty i \ \_ivl i = {}" -by(auto simp add: \_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 "\ = {\}" - -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 "\ = empty" - -instance -proof - case goal1 thus ?case - by (simp add:meet_ivl_def empty_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: 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 .. - -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 gamma_minus_ivl: - "n1 : \_ivl i1 \ n2 : \_ivl i2 \ n1-n2 : \_ivl(minus_ivl i1 i2)" -by(auto simp add: minus_ivl_def \_ivl_def split: ivl.splits option.splits) - -definition "filter_plus_ivl i i1 i2 = ((*if is_empty i then empty else*) - i1 \ minus_ivl i i2, i2 \ minus_ivl i i1)" - -fun filter_less_ivl :: "bool \ ivl \ ivl \ ivl * ivl" where -"filter_less_ivl res (I l1 h1) (I l2 h2) = - (if is_empty(I l1 h1) \ is_empty(I l2 h2) then (empty, empty) else - 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)))" - -global_interpretation Val_abs -where \ = \_ivl and num' = num_ivl and plus' = plus_ivl -proof - case goal1 thus ?case - by(auto simp: \_ivl_def le_ivl_def le_option_def split: ivl.split option.split if_splits) -next - case goal2 show ?case by(simp add: \_ivl_def Top_ivl_def) -next - case goal3 thus ?case by(simp add: \_ivl_def num_ivl_def) -next - case goal4 thus ?case - by(auto simp add: \_ivl_def plus_ivl_def split: ivl.split option.splits) -qed - -global_interpretation Val_abs1_gamma -where \ = \_ivl and num' = num_ivl and plus' = plus_ivl -defines aval_ivl = aval' -proof - case goal1 thus ?case - by(auto simp add: \_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 \_ivl_def empty_def) -qed - -lemma mono_minus_ivl: - "i1 \ i1' \ i2 \ i2' \ minus_ivl i1 i2 \ minus_ivl i1' i2'" -apply(auto simp add: minus_ivl_def empty_def le_ivl_def le_option_def split: ivl.splits) - apply(simp split: option.splits) - apply(simp split: option.splits) -apply(simp split: option.splits) -done - - -global_interpretation Val_abs1 -where \ = \_ivl and num' = num_ivl and plus' = plus_ivl -and test_num' = in_ivl -and filter_plus' = filter_plus_ivl and filter_less' = filter_less_ivl -proof - case goal1 thus ?case - by (simp add: \_ivl_def split: ivl.split option.split) -next - case goal2 thus ?case - by(auto simp add: filter_plus_ivl_def) - (metis gamma_minus_ivl add_diff_cancel add.commute)+ -next - case goal3 thus ?case - by(cases a1, cases a2, - auto simp: \_ivl_def min_option_def max_option_def le_option_def split: if_splits option.splits) -qed - -global_interpretation Abs_Int1 -where \ = \_ivl and num' = num_ivl and plus' = plus_ivl -and test_num' = in_ivl -and filter_plus' = filter_plus_ivl and filter_less' = filter_less_ivl -defines afilter_ivl = afilter -and bfilter_ivl = bfilter -and step_ivl = step' -and AI_ivl = AI -and aval_ivl' = aval'' -.. - - -text{* Monotonicity: *} - -global_interpretation Abs_Int1_mono -where \ = \_ivl and num' = num_ivl and plus' = plus_ivl -and test_num' = in_ivl -and filter_plus' = filter_plus_ivl and filter_less' = filter_less_ivl -proof - case goal1 thus ?case - by(auto simp: plus_ivl_def le_ivl_def le_option_def empty_def split: if_splits ivl.splits option.splits) -next - case goal2 thus ?case - by(auto simp: filter_plus_ivl_def le_prod_def mono_meet mono_minus_ivl) -next - case goal3 thus ?case - apply(cases a1, cases b1, cases a2, cases b2, auto simp: le_prod_def) - by(auto simp add: empty_def le_ivl_def le_option_def min_option_def max_option_def split: option.splits) -qed - - -subsubsection "Tests" - -value "show_acom_opt (AI_ivl test1_ivl)" - -text{* Better than @{text AI_const}: *} -value "show_acom_opt (AI_ivl test3_const)" -value "show_acom_opt (AI_ivl test4_const)" -value "show_acom_opt (AI_ivl test6_const)" - -value "show_acom_opt (AI_ivl test2_ivl)" -value "show_acom (((step_ivl \)^^0) (\\<^sub>c test2_ivl))" -value "show_acom (((step_ivl \)^^1) (\\<^sub>c test2_ivl))" -value "show_acom (((step_ivl \)^^2) (\\<^sub>c test2_ivl))" - -text{* Fixed point reached in 2 steps. Not so if the start value of x is known: *} - -value "show_acom_opt (AI_ivl test3_ivl)" -value "show_acom (((step_ivl \)^^0) (\\<^sub>c test3_ivl))" -value "show_acom (((step_ivl \)^^1) (\\<^sub>c test3_ivl))" -value "show_acom (((step_ivl \)^^2) (\\<^sub>c test3_ivl))" -value "show_acom (((step_ivl \)^^3) (\\<^sub>c test3_ivl))" -value "show_acom (((step_ivl \)^^4) (\\<^sub>c test3_ivl))" - -text{* Takes as many iterations as the actual execution. Would diverge if -loop did not terminate. Worse still, as the following example shows: even if -the actual execution terminates, the analysis may not. The value of y keeps -decreasing as the analysis is iterated, no matter how long: *} - -value "show_acom (((step_ivl \)^^50) (\\<^sub>c test4_ivl))" - -text{* Relationships between variables are NOT captured: *} -value "show_acom_opt (AI_ivl test5_ivl)" - -text{* Again, the analysis would not terminate: *} -value "show_acom (((step_ivl \)^^50) (\\<^sub>c test6_ivl))" - -end diff -r e44f5c123f26 -r 49b78f1f9e01 src/HOL/IMP/Abs_Int_ITP/Abs_Int3_ITP.thy --- a/src/HOL/IMP/Abs_Int_ITP/Abs_Int3_ITP.thy Tue Nov 22 18:36:59 2016 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,683 +0,0 @@ -(* Author: Tobias Nipkow *) - -theory Abs_Int3_ITP -imports Abs_Int2_ivl_ITP -begin - -subsection "Widening and Narrowing" - -class WN = SL_top + -fixes widen :: "'a \ 'a \ 'a" (infix "\" 65) -assumes widen1: "x \ x \ y" -assumes widen2: "y \ x \ y" -fixes narrow :: "'a \ 'a \ 'a" (infix "\" 65) -assumes narrow1: "y \ x \ y \ x \ y" -assumes narrow2: "y \ x \ x \ y \ x" - - -subsubsection "Intervals" - -instantiation ivl :: WN -begin - -definition "widen_ivl ivl1 ivl2 = - ((*if is_empty ivl1 then ivl2 else - if is_empty ivl2 then ivl1 else*) - case (ivl1,ivl2) of (I l1 h1, I l2 h2) \ - I (if le_option False l2 l1 \ l2 \ l1 then None else l1) - (if le_option True h1 h2 \ h1 \ h2 then None else h1))" - -definition "narrow_ivl ivl1 ivl2 = - ((*if is_empty ivl1 \ is_empty ivl2 then empty else*) - case (ivl1,ivl2) of (I l1 h1, I l2 h2) \ - I (if l1 = None then l2 else l1) - (if h1 = None then h2 else h1))" - -instance -proof qed - (auto simp add: widen_ivl_def narrow_ivl_def le_option_def le_ivl_def empty_def split: ivl.split option.split if_splits) - -end - - -subsubsection "Abstract State" - -instantiation st :: (WN)WN -begin - -definition "widen_st F1 F2 = - FunDom (\x. fun F1 x \ fun F2 x) (inter_list (dom F1) (dom F2))" - -definition "narrow_st F1 F2 = - FunDom (\x. fun F1 x \ fun F2 x) (inter_list (dom F1) (dom F2))" - -instance -proof - case goal1 thus ?case - by(simp add: widen_st_def le_st_def lookup_def widen1) -next - case goal2 thus ?case - by(simp add: widen_st_def le_st_def lookup_def widen2) -next - case goal3 thus ?case - by(auto simp: narrow_st_def le_st_def lookup_def narrow1) -next - case goal4 thus ?case - by(auto simp: narrow_st_def le_st_def lookup_def narrow2) -qed - -end - - -subsubsection "Option" - -instantiation option :: (WN)WN -begin - -fun widen_option where -"None \ x = x" | -"x \ None = x" | -"(Some x) \ (Some y) = Some(x \ y)" - -fun narrow_option where -"None \ x = None" | -"x \ None = None" | -"(Some x) \ (Some y) = Some(x \ y)" - -instance -proof - case goal1 show ?case - by(induct x y rule: widen_option.induct) (simp_all add: widen1) -next - case goal2 show ?case - by(induct x y rule: widen_option.induct) (simp_all add: widen2) -next - case goal3 thus ?case - by(induct x y rule: narrow_option.induct) (simp_all add: narrow1) -next - case goal4 thus ?case - by(induct x y rule: narrow_option.induct) (simp_all add: narrow2) -qed - -end - - -subsubsection "Annotated commands" - -fun map2_acom :: "('a \ 'a \ 'a) \ 'a acom \ 'a acom \ 'a acom" where -"map2_acom f (SKIP {a1}) (SKIP {a2}) = (SKIP {f a1 a2})" | -"map2_acom f (x ::= e {a1}) (x' ::= e' {a2}) = (x ::= e {f a1 a2})" | -"map2_acom f (c1;;c2) (c1';;c2') = (map2_acom f c1 c1';; map2_acom f c2 c2')" | -"map2_acom f (IF b THEN c1 ELSE c2 {a1}) (IF b' THEN c1' ELSE c2' {a2}) = - (IF b THEN map2_acom f c1 c1' ELSE map2_acom f c2 c2' {f a1 a2})" | -"map2_acom f ({a1} WHILE b DO c {a2}) ({a3} WHILE b' DO c' {a4}) = - ({f a1 a3} WHILE b DO map2_acom f c c' {f a2 a4})" - -abbreviation widen_acom :: "('a::WN)acom \ 'a acom \ 'a acom" (infix "\\<^sub>c" 65) -where "widen_acom == map2_acom (op \)" - -abbreviation narrow_acom :: "('a::WN)acom \ 'a acom \ 'a acom" (infix "\\<^sub>c" 65) -where "narrow_acom == map2_acom (op \)" - -lemma widen1_acom: "strip c = strip c' \ c \ c \\<^sub>c c'" -by(induct c c' rule: le_acom.induct)(simp_all add: widen1) - -lemma widen2_acom: "strip c = strip c' \ c' \ c \\<^sub>c c'" -by(induct c c' rule: le_acom.induct)(simp_all add: widen2) - -lemma narrow1_acom: "y \ x \ y \ x \\<^sub>c y" -by(induct y x rule: le_acom.induct) (simp_all add: narrow1) - -lemma narrow2_acom: "y \ x \ x \\<^sub>c y \ x" -by(induct y x rule: le_acom.induct) (simp_all add: narrow2) - - -subsubsection "Post-fixed point computation" - -definition iter_widen :: "('a acom \ 'a acom) \ 'a acom \ ('a::WN)acom option" -where "iter_widen f = while_option (\c. \ f c \ c) (\c. c \\<^sub>c f c)" - -definition iter_narrow :: "('a acom \ 'a acom) \ 'a acom \ 'a::WN acom option" -where "iter_narrow f = while_option (\c. \ c \ c \\<^sub>c f c) (\c. c \\<^sub>c f c)" - -definition pfp_wn :: - "(('a::WN)option acom \ 'a option acom) \ com \ 'a option acom option" -where "pfp_wn f c = (case iter_widen f (\\<^sub>c c) of None \ None - | Some c' \ iter_narrow f c')" - -lemma strip_map2_acom: - "strip c1 = strip c2 \ strip(map2_acom f c1 c2) = strip c1" -by(induct f c1 c2 rule: map2_acom.induct) simp_all - -lemma iter_widen_pfp: "iter_widen f c = Some c' \ f c' \ c'" -by(auto simp add: iter_widen_def dest: while_option_stop) - -lemma strip_while: fixes f :: "'a acom \ 'a acom" -assumes "\c. strip (f c) = strip c" and "while_option P f c = Some c'" -shows "strip c' = strip c" -using while_option_rule[where P = "\c'. strip c' = strip c", OF _ assms(2)] -by (metis assms(1)) - -lemma strip_iter_widen: fixes f :: "'a::WN acom \ 'a acom" -assumes "\c. strip (f c) = strip c" and "iter_widen f c = Some c'" -shows "strip c' = strip c" -proof- - have "\c. strip(c \\<^sub>c f c) = strip c" by (metis assms(1) strip_map2_acom) - from strip_while[OF this] assms(2) show ?thesis by(simp add: iter_widen_def) -qed - -lemma iter_narrow_pfp: assumes "mono f" and "f c0 \ c0" -and "iter_narrow f c0 = Some c" -shows "f c \ c \ c \ c0" (is "?P c") -proof- - { fix c assume "?P c" - note 1 = conjunct1[OF this] and 2 = conjunct2[OF this] - let ?c' = "c \\<^sub>c f c" - have "?P ?c'" - proof - have "f ?c' \ f c" by(rule monoD[OF `mono f` narrow2_acom[OF 1]]) - also have "\ \ ?c'" by(rule narrow1_acom[OF 1]) - finally show "f ?c' \ ?c'" . - have "?c' \ c" by (rule narrow2_acom[OF 1]) - also have "c \ c0" by(rule 2) - finally show "?c' \ c0" . - qed - } - with while_option_rule[where P = ?P, OF _ assms(3)[simplified iter_narrow_def]] - assms(2) le_refl - show ?thesis by blast -qed - -lemma pfp_wn_pfp: - "\ mono f; pfp_wn f c = Some c' \ \ f c' \ c'" -unfolding pfp_wn_def -by (auto dest: iter_widen_pfp iter_narrow_pfp split: option.splits) - -lemma strip_pfp_wn: - "\ \c. strip(f c) = strip c; pfp_wn f c = Some c' \ \ strip c' = c" -apply(auto simp add: pfp_wn_def iter_narrow_def split: option.splits) -by (metis (mono_tags) strip_map2_acom strip_while strip_bot_acom strip_iter_widen) - -locale Abs_Int2 = Abs_Int1_mono -where \=\ for \ :: "'av::{WN,L_top_bot} \ val set" -begin - -definition AI_wn :: "com \ 'av st option acom option" where -"AI_wn = pfp_wn (step' \)" - -lemma AI_wn_sound: "AI_wn c = Some c' \ CS c \ \\<^sub>c c'" -proof(simp add: CS_def AI_wn_def) - assume 1: "pfp_wn (step' \) c = Some c'" - from pfp_wn_pfp[OF mono_step'2 1] - have 2: "step' \ c' \ c'" . - have 3: "strip (\\<^sub>c (step' \ c')) = c" by(simp add: strip_pfp_wn[OF _ 1]) - have "lfp (step UNIV) c \ \\<^sub>c (step' \ c')" - proof(rule lfp_lowerbound[simplified,OF 3]) - show "step UNIV (\\<^sub>c (step' \ c')) \ \\<^sub>c (step' \ c')" - proof(rule step_preserves_le[OF _ _]) - show "UNIV \ \\<^sub>o \" by simp - show "\\<^sub>c (step' \ c') \ \\<^sub>c c'" by(rule mono_gamma_c[OF 2]) - qed - qed - from this 2 show "lfp (step UNIV) c \ \\<^sub>c c'" - by (blast intro: mono_gamma_c order_trans) -qed - -end - -global_interpretation Abs_Int2 -where \ = \_ivl and num' = num_ivl and plus' = plus_ivl -and test_num' = in_ivl -and filter_plus' = filter_plus_ivl and filter_less' = filter_less_ivl -defines AI_ivl' = AI_wn -.. - - -subsubsection "Tests" - -definition "step_up_ivl n = ((\c. c \\<^sub>c step_ivl \ c)^^n)" -definition "step_down_ivl n = ((\c. c \\<^sub>c step_ivl \ c)^^n)" - -text{* For @{const test3_ivl}, @{const AI_ivl} needed as many iterations as -the loop took to execute. In contrast, @{const AI_ivl'} converges in a -constant number of steps: *} - -value "show_acom (step_up_ivl 1 (\\<^sub>c test3_ivl))" -value "show_acom (step_up_ivl 2 (\\<^sub>c test3_ivl))" -value "show_acom (step_up_ivl 3 (\\<^sub>c test3_ivl))" -value "show_acom (step_up_ivl 4 (\\<^sub>c test3_ivl))" -value "show_acom (step_up_ivl 5 (\\<^sub>c test3_ivl))" -value "show_acom (step_down_ivl 1 (step_up_ivl 5 (\\<^sub>c test3_ivl)))" -value "show_acom (step_down_ivl 2 (step_up_ivl 5 (\\<^sub>c test3_ivl)))" -value "show_acom (step_down_ivl 3 (step_up_ivl 5 (\\<^sub>c test3_ivl)))" - -text{* Now all the analyses terminate: *} - -value "show_acom_opt (AI_ivl' test4_ivl)" -value "show_acom_opt (AI_ivl' test5_ivl)" -value "show_acom_opt (AI_ivl' test6_ivl)" - - -subsubsection "Termination: Intervals" - -definition m_ivl :: "ivl \ nat" where -"m_ivl ivl = (case ivl of I l h \ - (case l of None \ 0 | Some _ \ 1) + (case h of None \ 0 | Some _ \ 1))" - -lemma m_ivl_height: "m_ivl ivl \ 2" -by(simp add: m_ivl_def split: ivl.split option.split) - -lemma m_ivl_anti_mono: "(y::ivl) \ x \ m_ivl x \ m_ivl y" -by(auto simp: m_ivl_def le_option_def le_ivl_def - split: ivl.split option.split if_splits) - -lemma m_ivl_widen: - "~ y \ x \ m_ivl(x \ y) < m_ivl x" -by(auto simp: m_ivl_def widen_ivl_def le_option_def le_ivl_def - split: ivl.splits option.splits if_splits) - -lemma Top_less_ivl: "\ \ x \ m_ivl x = 0" -by(auto simp: m_ivl_def le_option_def le_ivl_def empty_def Top_ivl_def - split: ivl.split option.split if_splits) - - -definition n_ivl :: "ivl \ nat" where -"n_ivl ivl = 2 - m_ivl ivl" - -lemma n_ivl_mono: "(x::ivl) \ y \ n_ivl x \ n_ivl y" -unfolding n_ivl_def by (metis diff_le_mono2 m_ivl_anti_mono) - -lemma n_ivl_narrow: - "~ x \ x \ y \ n_ivl(x \ y) < n_ivl x" -by(auto simp: n_ivl_def m_ivl_def narrow_ivl_def le_option_def le_ivl_def - split: ivl.splits option.splits if_splits) - - -subsubsection "Termination: Abstract State" - -definition "m_st m st = (\x\set(dom st). m(fun st x))" - -lemma m_st_height: assumes "finite X" and "set (dom S) \ X" -shows "m_st m_ivl S \ 2 * card X" -proof(auto simp: m_st_def) - have "(\x\set(dom S). m_ivl (fun S x)) \ (\x\set(dom S). 2)" (is "?L \ _") - by(rule sum_mono)(simp add:m_ivl_height) - also have "\ \ (\x\X. 2)" - by(rule sum_mono3[OF assms]) simp - also have "\ = 2 * card X" by(simp add: sum_constant) - finally show "?L \ \" . -qed - -lemma m_st_anti_mono: - "S1 \ S2 \ m_st m_ivl S2 \ m_st m_ivl S1" -proof(auto simp: m_st_def le_st_def lookup_def split: if_splits) - let ?X = "set(dom S1)" let ?Y = "set(dom S2)" - let ?f = "fun S1" let ?g = "fun S2" - assume asm: "\x\?Y. (x \ ?X \ ?f x \ ?g x) \ (x \ ?X \ \ \ ?g x)" - hence 1: "\y\?Y\?X. m_ivl(?g y) \ m_ivl(?f y)" by(simp add: m_ivl_anti_mono) - have 0: "\x\?Y-?X. m_ivl(?g x) = 0" using asm by (auto simp: Top_less_ivl) - have "(\y\?Y. m_ivl(?g y)) = (\y\(?Y-?X) \ (?Y\?X). m_ivl(?g y))" - by (metis Un_Diff_Int) - also have "\ = (\y\?Y-?X. m_ivl(?g y)) + (\y\?Y\?X. m_ivl(?g y))" - by(subst sum.union_disjoint) auto - also have "(\y\?Y-?X. m_ivl(?g y)) = 0" using 0 by simp - also have "0 + (\y\?Y\?X. m_ivl(?g y)) = (\y\?Y\?X. m_ivl(?g y))" by simp - also have "\ \ (\y\?Y\?X. m_ivl(?f y))" - by(rule sum_mono)(simp add: 1) - also have "\ \ (\y\?X. m_ivl(?f y))" - by(simp add: sum_mono3[of "?X" "?Y Int ?X", OF _ Int_lower2]) - finally show "(\y\?Y. m_ivl(?g y)) \ (\x\?X. m_ivl(?f x))" - by (metis add_less_cancel_left) -qed - -lemma m_st_widen: -assumes "\ S2 \ S1" shows "m_st m_ivl (S1 \ S2) < m_st m_ivl S1" -proof- - { let ?X = "set(dom S1)" let ?Y = "set(dom S2)" - let ?f = "fun S1" let ?g = "fun S2" - fix x assume "x \ ?X" "\ lookup S2 x \ ?f x" - have "(\x\?X\?Y. m_ivl(?f x \ ?g x)) < (\x\?X. m_ivl(?f x))" (is "?L < ?R") - proof cases - assume "x : ?Y" - have "?L < (\x\?X\?Y. m_ivl(?f x))" - proof(rule sum_strict_mono_ex1, simp) - show "\x\?X\?Y. m_ivl(?f x \ ?g x) \ m_ivl (?f x)" - by (metis m_ivl_anti_mono widen1) - next - show "\x\?X\?Y. m_ivl(?f x \ ?g x) < m_ivl(?f x)" - using `x:?X` `x:?Y` `\ lookup S2 x \ ?f x` - by (metis IntI m_ivl_widen lookup_def) - qed - also have "\ \ ?R" by(simp add: sum_mono3[OF _ Int_lower1]) - finally show ?thesis . - next - assume "x ~: ?Y" - have "?L \ (\x\?X\?Y. m_ivl(?f x))" - proof(rule sum_mono, simp) - fix x assume "x:?X \ x:?Y" show "m_ivl(?f x \ ?g x) \ m_ivl (?f x)" - by (metis m_ivl_anti_mono widen1) - qed - also have "\ < m_ivl(?f x) + \" - using m_ivl_widen[OF `\ lookup S2 x \ ?f x`] - by (metis Nat.le_refl add_strict_increasing gr0I not_less0) - also have "\ = (\y\insert x (?X\?Y). m_ivl(?f y))" - using `x ~: ?Y` by simp - also have "\ \ (\x\?X. m_ivl(?f x))" - by(rule sum_mono3)(insert `x:?X`, auto) - finally show ?thesis . - qed - } with assms show ?thesis - by(auto simp: le_st_def widen_st_def m_st_def Int_def) -qed - -definition "n_st m X st = (\x\X. m(lookup st x))" - -lemma n_st_mono: assumes "set(dom S1) \ X" "set(dom S2) \ X" "S1 \ S2" -shows "n_st n_ivl X S1 \ n_st n_ivl X S2" -proof- - have "(\x\X. n_ivl(lookup S1 x)) \ (\x\X. n_ivl(lookup S2 x))" - apply(rule sum_mono) using assms - by(auto simp: le_st_def lookup_def n_ivl_mono split: if_splits) - thus ?thesis by(simp add: n_st_def) -qed - -lemma n_st_narrow: -assumes "finite X" and "set(dom S1) \ X" "set(dom S2) \ X" -and "S2 \ S1" "\ S1 \ S1 \ S2" -shows "n_st n_ivl X (S1 \ S2) < n_st n_ivl X S1" -proof- - have 1: "\x\X. n_ivl (lookup (S1 \ S2) x) \ n_ivl (lookup S1 x)" - using assms(2-4) - by(auto simp: le_st_def narrow_st_def lookup_def n_ivl_mono narrow2 - split: if_splits) - have 2: "\x\X. n_ivl (lookup (S1 \ S2) x) < n_ivl (lookup S1 x)" - using assms(2-5) - by(auto simp: le_st_def narrow_st_def lookup_def intro: n_ivl_narrow - split: if_splits) - have "(\x\X. n_ivl(lookup (S1 \ S2) x)) < (\x\X. n_ivl(lookup S1 x))" - apply(rule sum_strict_mono_ex1[OF `finite X`]) using 1 2 by blast+ - thus ?thesis by(simp add: n_st_def) -qed - - -subsubsection "Termination: Option" - -definition "m_o m n opt = (case opt of None \ n+1 | Some x \ m x)" - -lemma m_o_anti_mono: "finite X \ domo S2 \ X \ S1 \ S2 \ - m_o (m_st m_ivl) (2 * card X) S2 \ m_o (m_st m_ivl) (2 * card X) S1" -apply(induction S1 S2 rule: le_option.induct) -apply(auto simp: domo_def m_o_def m_st_anti_mono le_SucI m_st_height - split: option.splits) -done - -lemma m_o_widen: "\ finite X; domo S2 \ X; \ S2 \ S1 \ \ - m_o (m_st m_ivl) (2 * card X) (S1 \ S2) < m_o (m_st m_ivl) (2 * card X) S1" -by(auto simp: m_o_def domo_def m_st_height less_Suc_eq_le m_st_widen - split: option.split) - -definition "n_o n opt = (case opt of None \ 0 | Some x \ n x + 1)" - -lemma n_o_mono: "domo S1 \ X \ domo S2 \ X \ S1 \ S2 \ - n_o (n_st n_ivl X) S1 \ n_o (n_st n_ivl X) S2" -apply(induction S1 S2 rule: le_option.induct) -apply(auto simp: domo_def n_o_def n_st_mono - split: option.splits) -done - -lemma n_o_narrow: - "\ finite X; domo S1 \ X; domo S2 \ X; S2 \ S1; \ S1 \ S1 \ S2 \ - \ n_o (n_st n_ivl X) (S1 \ S2) < n_o (n_st n_ivl X) S1" -apply(induction S1 S2 rule: narrow_option.induct) -apply(auto simp: n_o_def domo_def n_st_narrow) -done - -lemma domo_widen_subset: "domo (S1 \ S2) \ domo S1 \ domo S2" -apply(induction S1 S2 rule: widen_option.induct) -apply (auto simp: domo_def widen_st_def) -done - -lemma domo_narrow_subset: "domo (S1 \ S2) \ domo S1 \ domo S2" -apply(induction S1 S2 rule: narrow_option.induct) -apply (auto simp: domo_def narrow_st_def) -done - -subsubsection "Termination: Commands" - -lemma strip_widen_acom[simp]: - "strip c' = strip (c::'a::WN acom) \ strip (c \\<^sub>c c') = strip c" -by(induction "widen::'a\'a\'a" c c' rule: map2_acom.induct) simp_all - -lemma strip_narrow_acom[simp]: - "strip c' = strip (c::'a::WN acom) \ strip (c \\<^sub>c c') = strip c" -by(induction "narrow::'a\'a\'a" c c' rule: map2_acom.induct) simp_all - -lemma annos_widen_acom[simp]: "strip c1 = strip (c2::'a::WN acom) \ - annos(c1 \\<^sub>c c2) = map (%(x,y).x\y) (zip (annos c1) (annos(c2::'a::WN acom)))" -by(induction "widen::'a\'a\'a" c1 c2 rule: map2_acom.induct) - (simp_all add: size_annos_same2) - -lemma annos_narrow_acom[simp]: "strip c1 = strip (c2::'a::WN acom) \ - annos(c1 \\<^sub>c c2) = map (%(x,y).x\y) (zip (annos c1) (annos(c2::'a::WN acom)))" -by(induction "narrow::'a\'a\'a" c1 c2 rule: map2_acom.induct) - (simp_all add: size_annos_same2) - -lemma widen_acom_Com[simp]: "strip c2 = strip c1 \ - c1 : Com X \ c2 : Com X \ (c1 \\<^sub>c c2) : Com X" -apply(auto simp add: Com_def) -apply(rename_tac S S' x) -apply(erule in_set_zipE) -apply(auto simp: domo_def split: option.splits) -apply(case_tac S) -apply(case_tac S') -apply simp -apply fastforce -apply(case_tac S') -apply fastforce -apply (fastforce simp: widen_st_def) -done - -lemma narrow_acom_Com[simp]: "strip c2 = strip c1 \ - c1 : Com X \ c2 : Com X \ (c1 \\<^sub>c c2) : Com X" -apply(auto simp add: Com_def) -apply(rename_tac S S' x) -apply(erule in_set_zipE) -apply(auto simp: domo_def split: option.splits) -apply(case_tac S) -apply(case_tac S') -apply simp -apply fastforce -apply(case_tac S') -apply fastforce -apply (fastforce simp: narrow_st_def) -done - -definition "m_c m c = (let as = annos c in \i=0.. {(c, c \\<^sub>c c') |c c'::ivl st option acom. - strip c' = strip c \ c : Com X \ c' : Com X \ \ c' \ c}\ - \ measure(m_c(m_o (m_st m_ivl) (2*card(X))))" -apply(auto simp: m_c_def Let_def Com_def) -apply(subgoal_tac "length(annos c') = length(annos c)") -prefer 2 apply (simp add: size_annos_same2) -apply (auto) -apply(rule sum_strict_mono_ex1) -apply simp -apply (clarsimp) -apply(erule m_o_anti_mono) -apply(rule subset_trans[OF domo_widen_subset]) -apply fastforce -apply(rule widen1) -apply(auto simp: le_iff_le_annos listrel_iff_nth) -apply(rule_tac x=n in bexI) -prefer 2 apply simp -apply(erule m_o_widen) -apply (simp)+ -done - -lemma measure_n_c: "finite X \ {(c, c \\<^sub>c c') |c c'. - strip c = strip c' \ c \ Com X \ c' \ Com X \ c' \ c \ \ c \ c \\<^sub>c c'}\ - \ measure(m_c(n_o (n_st n_ivl X)))" -apply(auto simp: m_c_def Let_def Com_def) -apply(subgoal_tac "length(annos c') = length(annos c)") -prefer 2 apply (simp add: size_annos_same2) -apply (auto) -apply(rule sum_strict_mono_ex1) -apply simp -apply (clarsimp) -apply(rule n_o_mono) -using domo_narrow_subset apply fastforce -apply fastforce -apply(rule narrow2) -apply(fastforce simp: le_iff_le_annos listrel_iff_nth) -apply(auto simp: le_iff_le_annos listrel_iff_nth strip_narrow_acom) -apply(rule_tac x=n in bexI) -prefer 2 apply simp -apply(erule n_o_narrow) -apply (simp)+ -done - - -subsubsection "Termination: Post-Fixed Point Iterations" - -lemma iter_widen_termination: -fixes c0 :: "'a::WN acom" -assumes P_f: "\c. P c \ P(f c)" -assumes P_widen: "\c c'. P c \ P c' \ P(c \\<^sub>c c')" -and "wf({(c::'a acom,c \\<^sub>c c')|c c'. P c \ P c' \ ~ c' \ c}^-1)" -and "P c0" and "c0 \ f c0" shows "EX c. iter_widen f c0 = Some c" -proof(simp add: iter_widen_def, rule wf_while_option_Some[where P = "P"]) - show "wf {(cc', c). (P c \ \ f c \ c) \ cc' = c \\<^sub>c f c}" - apply(rule wf_subset[OF assms(3)]) by(blast intro: P_f) -next - show "P c0" by(rule `P c0`) -next - fix c assume "P c" thus "P (c \\<^sub>c f c)" by(simp add: P_f P_widen) -qed - -lemma iter_narrow_termination: -assumes P_f: "\c. P c \ P(c \\<^sub>c f c)" -and wf: "wf({(c, c \\<^sub>c f c)|c c'. P c \ ~ c \ c \\<^sub>c f c}^-1)" -and "P c0" shows "EX c. iter_narrow f c0 = Some c" -proof(simp add: iter_narrow_def, rule wf_while_option_Some[where P = "P"]) - show "wf {(c', c). (P c \ \ c \ c \\<^sub>c f c) \ c' = c \\<^sub>c f c}" - apply(rule wf_subset[OF wf]) by(blast intro: P_f) -next - show "P c0" by(rule `P c0`) -next - fix c assume "P c" thus "P (c \\<^sub>c f c)" by(simp add: P_f) -qed - -lemma iter_winden_step_ivl_termination: - "EX c. iter_widen (step_ivl \) (\\<^sub>c c0) = Some c" -apply(rule iter_widen_termination[where - P = "%c. strip c = c0 \ c : Com(vars c0)"]) -apply (simp_all add: step'_Com bot_acom) -apply(rule wf_subset) -apply(rule wf_measure) -apply(rule subset_trans) -prefer 2 -apply(rule measure_m_c[where X = "vars c0", OF finite_cvars]) -apply blast -done - -lemma iter_narrow_step_ivl_termination: - "c0 \ Com (vars(strip c0)) \ step_ivl \ c0 \ c0 \ - EX c. iter_narrow (step_ivl \) c0 = Some c" -apply(rule iter_narrow_termination[where - P = "%c. strip c = strip c0 \ c : Com(vars(strip c0)) \ step_ivl \ c \ c"]) -apply (simp_all add: step'_Com) -apply(clarify) -apply(frule narrow2_acom, drule mono_step'[OF le_refl], erule le_trans[OF _ narrow1_acom]) -apply assumption -apply(rule wf_subset) -apply(rule wf_measure) -apply(rule subset_trans) -prefer 2 -apply(rule measure_n_c[where X = "vars(strip c0)", OF finite_cvars]) -apply auto -by (metis bot_least domo_Top order_refl step'_Com strip_step') - -(* FIXME: simplify type system: Combine Com(X) and vars <= X?? *) -lemma while_Com: -fixes c :: "'a st option acom" -assumes "while_option P f c = Some c'" -and "!!c. strip(f c) = strip c" -and "\c::'a st option acom. c : Com(X) \ vars(strip c) \ X \ f c : Com(X)" -and "c : Com(X)" and "vars(strip c) \ X" shows "c' : Com(X)" -using while_option_rule[where P = "\c'. c' : Com(X) \ vars(strip c') \ X", OF _ assms(1)] -by(simp add: assms(2-)) - -lemma iter_widen_Com: fixes f :: "'a::WN st option acom \ 'a st option acom" -assumes "iter_widen f c = Some c'" -and "\c. c : Com(X) \ vars(strip c) \ X \ f c : Com(X)" -and "!!c. strip(f c) = strip c" -and "c : Com(X)" and "vars (strip c) \ X" shows "c' : Com(X)" -proof- - have "\c. c : Com(X) \ vars(strip c) \ X \ c \\<^sub>c f c : Com(X)" - by (metis (full_types) widen_acom_Com assms(2,3)) - from while_Com[OF assms(1)[simplified iter_widen_def] _ this assms(4,5)] - show ?thesis using assms(3) by(simp) -qed - - -context Abs_Int2 -begin - -lemma iter_widen_step'_Com: - "iter_widen (step' \) c = Some c' \ vars(strip c) \ X \ c : Com(X) - \ c' : Com(X)" -apply(subgoal_tac "strip c'= strip c") -prefer 2 apply (metis strip_iter_widen strip_step') -apply(drule iter_widen_Com) -prefer 3 apply assumption -prefer 3 apply assumption -apply (auto simp: step'_Com) -done - -end - -theorem AI_ivl'_termination: - "EX c'. AI_ivl' c = Some c'" -apply(auto simp: AI_wn_def pfp_wn_def iter_winden_step_ivl_termination split: option.split) -apply(rule iter_narrow_step_ivl_termination) -apply (metis bot_acom_Com iter_widen_step'_Com[OF _ subset_refl] strip_iter_widen strip_step') -apply(erule iter_widen_pfp) -done - -end - - -(* interesting(?) relic -lemma widen_assoc: - "~ (y::ivl) \ x \ ~ z \ x \ y \ ((x::ivl) \ y) \ z = x \ (y \ z)" -apply(cases x) -apply(cases y) -apply(cases z) -apply(rename_tac x1 x2 y1 y2 z1 z2) -apply(simp add: le_ivl_def) -apply(case_tac x1) -apply(case_tac x2) -apply(simp add:le_option_def widen_ivl_def split: if_splits option.splits) -apply(simp add:le_option_def widen_ivl_def split: if_splits option.splits) -apply(case_tac x2) -apply(simp add:le_option_def widen_ivl_def split: if_splits option.splits) -apply(case_tac y1) -apply(case_tac y2) -apply(simp add:le_option_def widen_ivl_def split: if_splits option.splits) -apply(case_tac z1) -apply(auto simp add:le_option_def widen_ivl_def split: if_splits option.splits ivl.splits)[1] -apply(auto simp add:le_option_def widen_ivl_def split: if_splits option.splits ivl.splits)[1] -apply(case_tac y2) -apply(auto simp add:le_option_def widen_ivl_def split: if_splits option.splits ivl.splits)[1] -apply(case_tac z1) -apply(auto simp add:le_option_def widen_ivl_def split: if_splits ivl.splits option.splits)[1] -apply(case_tac z2) -apply(auto simp add:le_option_def widen_ivl_def split: if_splits option.splits)[1] -apply(auto simp add:le_option_def widen_ivl_def split: if_splits option.splits)[1] -done - -lemma widen_step_trans: - "~ (y::ivl) \ x \ ~ z \ x \ y \ EX u. (x \ y) \ z = x \ u \ ~ u \ x" -by (metis widen_assoc preord_class.le_trans widen1) -*) diff -r e44f5c123f26 -r 49b78f1f9e01 src/HOL/IMP/Abs_Int_ITP/Abs_State_ITP.thy --- a/src/HOL/IMP/Abs_Int_ITP/Abs_State_ITP.thy Tue Nov 22 18:36:59 2016 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,98 +0,0 @@ -(* Author: Tobias Nipkow *) - -theory Abs_State_ITP -imports Abs_Int0_ITP - "~~/src/HOL/Library/Char_ord" "~~/src/HOL/Library/List_lexord" - (* Library import merely to allow string lists to be sorted for output *) -begin - -subsection "Abstract State with Computable Ordering" - -text{* A concrete type of state with computable @{text"\"}: *} - -datatype 'a st = FunDom "vname \ 'a" "vname list" - -fun "fun" where "fun (FunDom f xs) = f" -fun dom where "dom (FunDom f xs) = xs" - -definition [simp]: "inter_list xs ys = [x\xs. x \ set ys]" - -definition "show_st S = [(x,fun S x). x \ sort(dom S)]" - -definition "show_acom = map_acom (map_option show_st)" -definition "show_acom_opt = map_option show_acom" - -definition "lookup F x = (if x : set(dom F) then fun F x else \)" - -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) - -definition "\_st \ F = {f. \x. f x \ \(lookup F x)}" - -instantiation st :: (SL_top) SL_top -begin - -definition "le_st F G = (ALL x : set(dom G). lookup F x \ fun G x)" - -definition -"join_st F G = - FunDom (\x. fun F x \ fun G x) (inter_list (dom F) (dom G))" - -definition "\ = FunDom (\x. \) []" - -instance -proof - case goal2 thus ?case - apply(auto simp: le_st_def) - by (metis lookup_def preord_class.le_trans top) -qed (auto simp: le_st_def lookup_def join_st_def Top_st_def) - -end - -lemma mono_lookup: "F \ F' \ lookup F x \ lookup F' x" -by(auto simp add: lookup_def le_st_def) - -lemma mono_update: "a \ a' \ S \ S' \ update S x a \ update S' x a'" -by(auto simp add: le_st_def lookup_def update_def) - -locale Gamma = Val_abs where \=\ for \ :: "'av::SL_top \ val set" -begin - -abbreviation \\<^sub>f :: "'av st \ state set" -where "\\<^sub>f == \_st \" - -abbreviation \\<^sub>o :: "'av st option \ state set" -where "\\<^sub>o == \_option \\<^sub>f" - -abbreviation \\<^sub>c :: "'av st option acom \ state set acom" -where "\\<^sub>c == map_acom \\<^sub>o" - -lemma gamma_f_Top[simp]: "\\<^sub>f Top = UNIV" -by(auto simp: Top_st_def \_st_def lookup_def) - -lemma gamma_o_Top[simp]: "\\<^sub>o Top = UNIV" -by (simp add: Top_option_def) - -(* FIXME (maybe also le \ sqle?) *) - -lemma mono_gamma_f: "f \ g \ \\<^sub>f f \ \\<^sub>f g" -apply(simp add:\_st_def subset_iff lookup_def le_st_def split: if_splits) -by (metis UNIV_I mono_gamma gamma_Top subsetD) - -lemma mono_gamma_o: - "sa \ sa' \ \\<^sub>o sa \ \\<^sub>o sa'" -by(induction sa sa' rule: le_option.induct)(simp_all add: mono_gamma_f) - -lemma mono_gamma_c: "ca \ ca' \ \\<^sub>c ca \ \\<^sub>c ca'" -by (induction ca ca' 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 diff -r e44f5c123f26 -r 49b78f1f9e01 src/HOL/IMP/Abs_Int_ITP/Collecting_ITP.thy --- a/src/HOL/IMP/Abs_Int_ITP/Collecting_ITP.thy Tue Nov 22 18:36:59 2016 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,183 +0,0 @@ -theory Collecting_ITP -imports Complete_Lattice_ix "ACom_ITP" -begin - -subsection "Collecting Semantics of Commands" - -subsubsection "Annotated commands as a complete lattice" - -(* Orderings could also be lifted generically (thus subsuming the -instantiation for preord and order), but then less_eq_acom would need to -become a definition, eg less_eq_acom = lift2 less_eq, and then proofs would -need to unfold this defn first. *) - -instantiation acom :: (order) order -begin - -fun less_eq_acom :: "('a::order)acom \ 'a acom \ bool" where -"(SKIP {S}) \ (SKIP {S'}) = (S \ S')" | -"(x ::= e {S}) \ (x' ::= e' {S'}) = (x=x' \ e=e' \ S \ S')" | -"(c1;;c2) \ (c1';;c2') = (c1 \ c1' \ c2 \ c2')" | -"(IF b THEN c1 ELSE c2 {S}) \ (IF b' THEN c1' ELSE c2' {S'}) = - (b=b' \ c1 \ c1' \ c2 \ c2' \ S \ S')" | -"({Inv} WHILE b DO c {P}) \ ({Inv'} WHILE b' DO c' {P'}) = - (b=b' \ c \ c' \ Inv \ Inv' \ P \ P')" | -"less_eq_acom _ _ = False" - -lemma SKIP_le: "SKIP {S} \ c \ (\S'. c = SKIP {S'} \ S \ S')" -by (cases c) auto - -lemma Assign_le: "x ::= e {S} \ c \ (\S'. c = x ::= e {S'} \ S \ S')" -by (cases c) auto - -lemma Seq_le: "c1;;c2 \ c \ (\c1' c2'. c = c1';;c2' \ c1 \ c1' \ c2 \ c2')" -by (cases c) auto - -lemma If_le: "IF b THEN c1 ELSE c2 {S} \ c \ - (\c1' c2' S'. c= IF b THEN c1' ELSE c2' {S'} \ c1 \ c1' \ c2 \ c2' \ S \ S')" -by (cases c) auto - -lemma While_le: "{Inv} WHILE b DO c {P} \ w \ - (\Inv' c' P'. w = {Inv'} WHILE b DO c' {P'} \ c \ c' \ Inv \ Inv' \ P \ P')" -by (cases w) auto - -definition less_acom :: "'a acom \ 'a acom \ bool" where -"less_acom x y = (x \ y \ \ y \ x)" - -instance -proof - case goal1 show ?case by(simp add: less_acom_def) -next - case goal2 thus ?case by (induct x) auto -next - case goal3 thus ?case - apply(induct x y arbitrary: z rule: less_eq_acom.induct) - apply (auto intro: le_trans simp: SKIP_le Assign_le Seq_le If_le While_le) - done -next - case goal4 thus ?case - apply(induct x y rule: less_eq_acom.induct) - apply (auto intro: le_antisym) - done -qed - -end - -fun sub\<^sub>1 :: "'a acom \ 'a acom" where -"sub\<^sub>1(c1;;c2) = c1" | -"sub\<^sub>1(IF b THEN c1 ELSE c2 {S}) = c1" | -"sub\<^sub>1({I} WHILE b DO c {P}) = c" - -fun sub\<^sub>2 :: "'a acom \ 'a acom" where -"sub\<^sub>2(c1;;c2) = c2" | -"sub\<^sub>2(IF b THEN c1 ELSE c2 {S}) = c2" - -fun invar :: "'a acom \ 'a" where -"invar({I} WHILE b DO c {P}) = I" - -fun lift :: "('a set \ 'b) \ com \ 'a acom set \ 'b acom" -where -"lift F com.SKIP M = (SKIP {F (post ` M)})" | -"lift F (x ::= a) M = (x ::= a {F (post ` M)})" | -"lift F (c1;;c2) M = - lift F c1 (sub\<^sub>1 ` M);; lift F c2 (sub\<^sub>2 ` M)" | -"lift F (IF b THEN c1 ELSE c2) M = - IF b THEN lift F c1 (sub\<^sub>1 ` M) ELSE lift F c2 (sub\<^sub>2 ` M) - {F (post ` M)}" | -"lift F (WHILE b DO c) M = - {F (invar ` M)} - WHILE b DO lift F c (sub\<^sub>1 ` M) - {F (post ` M)}" - -global_interpretation Complete_Lattice_ix "%c. {c'. strip c' = c}" "lift Inter" -proof - case goal1 - have "a:A \ lift Inter (strip a) A \ a" - proof(induction a arbitrary: A) - case Seq from Seq.prems show ?case by(force intro!: Seq.IH) - next - case If from If.prems show ?case by(force intro!: If.IH) - next - case While from While.prems show ?case by(force intro!: While.IH) - qed force+ - with goal1 show ?case by auto -next - case goal2 - thus ?case - proof(induction b arbitrary: i A) - case SKIP thus ?case by (force simp:SKIP_le) - next - case Assign thus ?case by (force simp:Assign_le) - next - case Seq from Seq.prems show ?case - by (force intro!: Seq.IH simp:Seq_le) - next - case If from If.prems show ?case by (force simp: If_le intro!: If.IH) - next - case While from While.prems show ?case - by(fastforce simp: While_le intro: While.IH) - qed -next - case goal3 - have "strip(lift Inter i A) = i" - proof(induction i arbitrary: A) - case Seq from Seq.prems show ?case - by (fastforce simp: strip_eq_Seq subset_iff intro!: Seq.IH) - next - case If from If.prems show ?case - by (fastforce intro!: If.IH simp: strip_eq_If) - next - case While from While.prems show ?case - by(fastforce intro: While.IH simp: strip_eq_While) - qed auto - thus ?case by auto -qed - -lemma le_post: "c \ d \ post c \ post d" -by(induction c d rule: less_eq_acom.induct) auto - -subsubsection "Collecting semantics" - -fun step :: "state set \ state set acom \ state set acom" where -"step S (SKIP {P}) = (SKIP {S})" | -"step S (x ::= e {P}) = - (x ::= e {{s'. EX s:S. s' = s(x := aval e s)}})" | -"step S (c1;; c2) = step S c1;; step (post c1) c2" | -"step S (IF b THEN c1 ELSE c2 {P}) = - IF b THEN step {s:S. bval b s} c1 ELSE step {s:S. \ bval b s} c2 - {post c1 \ post c2}" | -"step S ({Inv} WHILE b DO c {P}) = - {S \ post c} WHILE b DO (step {s:Inv. bval b s} c) {{s:Inv. \ bval b s}}" - -definition CS :: "com \ state set acom" where -"CS c = lfp (step UNIV) c" - -lemma mono2_step: "c1 \ c2 \ S1 \ S2 \ step S1 c1 \ step S2 c2" -proof(induction c1 c2 arbitrary: S1 S2 rule: less_eq_acom.induct) - case 2 thus ?case by fastforce -next - case 3 thus ?case by(simp add: le_post) -next - case 4 thus ?case by(simp add: subset_iff)(metis le_post set_mp)+ -next - case 5 thus ?case by(simp add: subset_iff) (metis le_post set_mp) -qed auto - -lemma mono_step: "mono (step S)" -by(blast intro: monoI mono2_step) - -lemma strip_step: "strip(step S c) = strip c" -by (induction c arbitrary: S) auto - -lemma lfp_cs_unfold: "lfp (step S) c = step S (lfp (step S) c)" -apply(rule lfp_unfold[OF _ mono_step]) -apply(simp add: strip_step) -done - -lemma CS_unfold: "CS c = step UNIV (CS c)" -by (metis CS_def lfp_cs_unfold) - -lemma strip_CS[simp]: "strip(CS c) = c" -by(simp add: CS_def index_lfp[simplified]) - -end diff -r e44f5c123f26 -r 49b78f1f9e01 src/HOL/IMP/Abs_Int_ITP/Complete_Lattice_ix.thy --- a/src/HOL/IMP/Abs_Int_ITP/Complete_Lattice_ix.thy Tue Nov 22 18:36:59 2016 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,61 +0,0 @@ -(* Author: Tobias Nipkow *) - -section "Abstract Interpretation (ITP)" - -theory Complete_Lattice_ix -imports Main -begin - -subsection "Complete Lattice (indexed)" - -text{* A complete lattice is an ordered type where every set of elements has -a greatest lower (and thus also a leats upper) bound. Sets are the -prototypical complete lattice where the greatest lower bound is -intersection. Sometimes that set of all elements of a type is not a complete -lattice although all elements of the same shape form a complete lattice, for -example lists of the same length, where the list elements come from a -complete lattice. We will have exactly this situation with annotated -commands. This theory introduces a slightly generalised version of complete -lattices where elements have an ``index'' and only the set of elements with -the same index form a complete lattice; the type as a whole is a disjoint -union of complete lattices. Because sets are not types, this requires a -special treatment. *} - -locale Complete_Lattice_ix = -fixes L :: "'i \ 'a::order set" -and Glb :: "'i \ 'a set \ 'a" -assumes Glb_lower: "A \ L i \ a \ A \ (Glb i A) \ a" -and Glb_greatest: "b : L i \ \a\A. b \ a \ b \ (Glb i A)" -and Glb_in_L: "A \ L i \ Glb i A : L i" -begin - -definition lfp :: "('a \ 'a) \ 'i \ 'a" where -"lfp f i = Glb i {a : L i. f a \ a}" - -lemma index_lfp: "lfp f i : L i" -by(auto simp: lfp_def intro: Glb_in_L) - -lemma lfp_lowerbound: - "\ a : L i; f a \ a \ \ lfp f i \ a" -by (auto simp add: lfp_def intro: Glb_lower) - -lemma lfp_greatest: - "\ a : L i; \u. \ u : L i; f u \ u\ \ a \ u \ \ a \ lfp f i" -by (auto simp add: lfp_def intro: Glb_greatest) - -lemma lfp_unfold: assumes "\x i. f x : L i \ x : L i" -and mono: "mono f" shows "lfp f i = f (lfp f i)" -proof- - note assms(1)[simp] index_lfp[simp] - have 1: "f (lfp f i) \ lfp f i" - apply(rule lfp_greatest) - apply simp - by (blast intro: lfp_lowerbound monoD[OF mono] order_trans) - have "lfp f i \ f (lfp f i)" - by (fastforce intro: 1 monoD[OF mono] lfp_lowerbound) - with 1 show ?thesis by(blast intro: order_antisym) -qed - -end - -end diff -r e44f5c123f26 -r 49b78f1f9e01 src/HOL/ROOT --- a/src/HOL/ROOT Tue Nov 22 18:36:59 2016 +0100 +++ b/src/HOL/ROOT Wed Nov 23 16:28:42 2016 +0100 @@ -153,9 +153,6 @@ Abs_Int1_parity Abs_Int1_const Abs_Int3 - "Abs_Int_ITP/Abs_Int1_parity_ITP" - "Abs_Int_ITP/Abs_Int1_const_ITP" - "Abs_Int_ITP/Abs_Int3_ITP" Procs_Dyn_Vars_Dyn Procs_Stat_Vars_Dyn Procs_Stat_Vars_Stat