| author | Andreas Lochbihler | 
| Fri, 19 Sep 2014 08:26:03 +0200 | |
| changeset 58384 | 00aaaa7bd752 | 
| parent 58310 | 91ea607a34d8 | 
| child 61671 | 20d4cd2ceab2 | 
| permissions | -rw-r--r-- | 
| 45111 | 1 | (* Author: Tobias Nipkow *) | 
| 2 | ||
| 47602 | 3 | theory Abs_Int1_const_ITP | 
| 48480 | 4 | imports Abs_Int1_ITP "../Abs_Int_Tests" | 
| 45111 | 5 | begin | 
| 6 | ||
| 7 | subsection "Constant Propagation" | |
| 8 | ||
| 58310 | 9 | datatype const = Const val | Any | 
| 45111 | 10 | |
| 46158 | 11 | fun \<gamma>_const where | 
| 12 | "\<gamma>_const (Const n) = {n}" |
 | |
| 13 | "\<gamma>_const (Any) = UNIV" | |
| 45111 | 14 | |
| 46158 | 15 | fun plus_const where | 
| 16 | "plus_const (Const m) (Const n) = Const(m+n)" | | |
| 17 | "plus_const _ _ = Any" | |
| 45111 | 18 | |
| 46158 | 19 | lemma plus_const_cases: "plus_const a1 a2 = | 
| 45111 | 20 | (case (a1,a2) of (Const m, Const n) \<Rightarrow> Const(m+n) | _ \<Rightarrow> Any)" | 
| 46158 | 21 | by(auto split: prod.split const.split) | 
| 45111 | 22 | |
| 46158 | 23 | instantiation const :: SL_top | 
| 45111 | 24 | begin | 
| 25 | ||
| 46158 | 26 | fun le_const where | 
| 45111 | 27 | "_ \<sqsubseteq> Any = True" | | 
| 28 | "Const n \<sqsubseteq> Const m = (n=m)" | | |
| 29 | "Any \<sqsubseteq> Const _ = False" | |
| 30 | ||
| 46158 | 31 | fun join_const where | 
| 45111 | 32 | "Const m \<squnion> Const n = (if n=m then Const m else Any)" | | 
| 33 | "_ \<squnion> _ = Any" | |
| 34 | ||
| 35 | definition "\<top> = Any" | |
| 36 | ||
| 37 | instance | |
| 38 | proof | |
| 39 | case goal1 thus ?case by (cases x) simp_all | |
| 40 | next | |
| 41 | case goal2 thus ?case by(cases z, cases y, cases x, simp_all) | |
| 42 | next | |
| 43 | case goal3 thus ?case by(cases x, cases y, simp_all) | |
| 44 | next | |
| 45 | case goal4 thus ?case by(cases y, cases x, simp_all) | |
| 46 | next | |
| 47 | case goal5 thus ?case by(cases z, cases y, cases x, simp_all) | |
| 48 | next | |
| 46158 | 49 | case goal6 thus ?case by(simp add: Top_const_def) | 
| 45111 | 50 | qed | 
| 51 | ||
| 52 | end | |
| 53 | ||
| 45127 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 nipkow parents: 
45111diff
changeset | 54 | |
| 55599 
6535c537b243
aggiornamento for "interpretation with definitions": operate uniformly on theory and locale level under the brand of "permanent interpretation"
 haftmann parents: 
48480diff
changeset | 55 | permanent_interpretation Val_abs | 
| 46158 | 56 | where \<gamma> = \<gamma>_const and num' = Const and plus' = plus_const | 
| 45111 | 57 | proof | 
| 58 | case goal1 thus ?case | |
| 59 | by(cases a, cases b, simp, simp, cases b, simp, simp) | |
| 60 | next | |
| 46158 | 61 | case goal2 show ?case by(simp add: Top_const_def) | 
| 45111 | 62 | next | 
| 45127 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 nipkow parents: 
45111diff
changeset | 63 | case goal3 show ?case by simp | 
| 45111 | 64 | next | 
| 45127 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 nipkow parents: 
45111diff
changeset | 65 | case goal4 thus ?case | 
| 46158 | 66 | by(auto simp: plus_const_cases split: const.split) | 
| 45111 | 67 | qed | 
| 68 | ||
| 55599 
6535c537b243
aggiornamento for "interpretation with definitions": operate uniformly on theory and locale level under the brand of "permanent interpretation"
 haftmann parents: 
48480diff
changeset | 69 | permanent_interpretation Abs_Int | 
| 46158 | 70 | where \<gamma> = \<gamma>_const and num' = Const and plus' = plus_const | 
| 55600 
3c7610b8dcfc
more convenient syntax for permanent interpretation
 haftmann parents: 
55599diff
changeset | 71 | defining AI_const = AI and step_const = step' and aval'_const = aval' | 
| 46355 
42a01315d998
removed accidental dependance of abstract interpreter on gamma
 nipkow parents: 
46353diff
changeset | 72 | .. | 
| 45111 | 73 | |
| 74 | ||
| 46353 | 75 | subsubsection "Tests" | 
| 76 | ||
| 77 | value "show_acom (((step_const \<top>)^^0) (\<bottom>\<^sub>c test1_const))" | |
| 78 | value "show_acom (((step_const \<top>)^^1) (\<bottom>\<^sub>c test1_const))" | |
| 79 | value "show_acom (((step_const \<top>)^^2) (\<bottom>\<^sub>c test1_const))" | |
| 80 | value "show_acom (((step_const \<top>)^^3) (\<bottom>\<^sub>c test1_const))" | |
| 81 | value "show_acom_opt (AI_const test1_const)" | |
| 82 | ||
| 83 | value "show_acom_opt (AI_const test2_const)" | |
| 84 | value "show_acom_opt (AI_const test3_const)" | |
| 85 | ||
| 86 | value "show_acom (((step_const \<top>)^^0) (\<bottom>\<^sub>c test4_const))" | |
| 87 | value "show_acom (((step_const \<top>)^^1) (\<bottom>\<^sub>c test4_const))" | |
| 88 | value "show_acom (((step_const \<top>)^^2) (\<bottom>\<^sub>c test4_const))" | |
| 89 | value "show_acom (((step_const \<top>)^^3) (\<bottom>\<^sub>c test4_const))" | |
| 90 | value "show_acom_opt (AI_const test4_const)" | |
| 91 | ||
| 92 | value "show_acom (((step_const \<top>)^^0) (\<bottom>\<^sub>c test5_const))" | |
| 93 | value "show_acom (((step_const \<top>)^^1) (\<bottom>\<^sub>c test5_const))" | |
| 94 | value "show_acom (((step_const \<top>)^^2) (\<bottom>\<^sub>c test5_const))" | |
| 95 | value "show_acom (((step_const \<top>)^^3) (\<bottom>\<^sub>c test5_const))" | |
| 96 | value "show_acom (((step_const \<top>)^^4) (\<bottom>\<^sub>c test5_const))" | |
| 97 | value "show_acom (((step_const \<top>)^^5) (\<bottom>\<^sub>c test5_const))" | |
| 98 | value "show_acom_opt (AI_const test5_const)" | |
| 99 | ||
| 100 | value "show_acom (((step_const \<top>)^^0) (\<bottom>\<^sub>c test6_const))" | |
| 101 | value "show_acom (((step_const \<top>)^^1) (\<bottom>\<^sub>c test6_const))" | |
| 102 | value "show_acom (((step_const \<top>)^^2) (\<bottom>\<^sub>c test6_const))" | |
| 103 | value "show_acom (((step_const \<top>)^^3) (\<bottom>\<^sub>c test6_const))" | |
| 104 | value "show_acom (((step_const \<top>)^^4) (\<bottom>\<^sub>c test6_const))" | |
| 105 | value "show_acom (((step_const \<top>)^^5) (\<bottom>\<^sub>c test6_const))" | |
| 106 | value "show_acom (((step_const \<top>)^^6) (\<bottom>\<^sub>c test6_const))" | |
| 107 | value "show_acom (((step_const \<top>)^^7) (\<bottom>\<^sub>c test6_const))" | |
| 108 | value "show_acom (((step_const \<top>)^^8) (\<bottom>\<^sub>c test6_const))" | |
| 109 | value "show_acom (((step_const \<top>)^^9) (\<bottom>\<^sub>c test6_const))" | |
| 110 | value "show_acom (((step_const \<top>)^^10) (\<bottom>\<^sub>c test6_const))" | |
| 111 | value "show_acom (((step_const \<top>)^^11) (\<bottom>\<^sub>c test6_const))" | |
| 112 | value "show_acom_opt (AI_const test6_const)" | |
| 113 | ||
| 114 | ||
| 45623 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: 
45200diff
changeset | 115 | text{* Monotonicity: *}
 | 
| 45111 | 116 | |
| 55599 
6535c537b243
aggiornamento for "interpretation with definitions": operate uniformly on theory and locale level under the brand of "permanent interpretation"
 haftmann parents: 
48480diff
changeset | 117 | permanent_interpretation Abs_Int_mono | 
| 46158 | 118 | where \<gamma> = \<gamma>_const and num' = Const and plus' = plus_const | 
| 45623 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: 
45200diff
changeset | 119 | proof | 
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: 
45200diff
changeset | 120 | case goal1 thus ?case | 
| 46158 | 121 | by(auto simp: plus_const_cases split: const.split) | 
| 45623 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: 
45200diff
changeset | 122 | qed | 
| 45111 | 123 | |
| 46158 | 124 | text{* Termination: *}
 | 
| 125 | ||
| 46334 | 126 | definition "m_const x = (case x of Const _ \<Rightarrow> 1 | Any \<Rightarrow> 0)" | 
| 127 | ||
| 46158 | 128 | lemma measure_const: | 
| 46334 | 129 |   "(strict{(x::const,y). x \<sqsubseteq> y})^-1 \<subseteq> measure m_const"
 | 
| 130 | by(auto simp: m_const_def split: const.splits) | |
| 46158 | 131 | |
| 132 | lemma measure_const_eq: | |
| 46334 | 133 | "\<forall> x y::const. x \<sqsubseteq> y \<and> y \<sqsubseteq> x \<longrightarrow> m_const x = m_const y" | 
| 134 | by(auto simp: m_const_def split: const.splits) | |
| 46158 | 135 | |
| 136 | lemma "EX c'. AI_const c = Some c'" | |
| 46334 | 137 | by(rule AI_Some_measure[OF measure_const measure_const_eq]) | 
| 46158 | 138 | |
| 45111 | 139 | end |