Theory Abs_Int1_const
theory Abs_Int1_const
imports Abs_Int1 "HOL-IMP.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 (standard, goal_cases)
case (1 x) thus ?case by (cases x) simp_all
next
case (2 x y z) thus ?case by(cases z, cases y, cases x, simp_all)
next
case (3 x y) thus ?case by(cases x, cases y, simp_all)
next
case (4 x y) thus ?case by(cases y, cases x, simp_all)
next
case (5 x y z) thus ?case by(cases z, cases y, cases x, simp_all)
next
case 6 thus ?case by(simp add: Top_const_def)
qed
end
global_interpretation Val_abs
where γ = γ_const and num' = Const and plus' = plus_const
proof (standard, goal_cases)
case (1 a b) thus ?case
by(cases a, cases b, simp, simp, cases b, simp, simp)
next
case 2 show ?case by(simp add: Top_const_def)
next
case 3 show ?case by simp
next
case 4 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) (⊥⇩c test1_const))"
value "show_acom (((step_const ⊤)^^1) (⊥⇩c test1_const))"
value "show_acom (((step_const ⊤)^^2) (⊥⇩c test1_const))"
value "show_acom (((step_const ⊤)^^3) (⊥⇩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) (⊥⇩c test4_const))"
value "show_acom (((step_const ⊤)^^1) (⊥⇩c test4_const))"
value "show_acom (((step_const ⊤)^^2) (⊥⇩c test4_const))"
value "show_acom (((step_const ⊤)^^3) (⊥⇩c test4_const))"
value "show_acom_opt (AI_const test4_const)"
value "show_acom (((step_const ⊤)^^0) (⊥⇩c test5_const))"
value "show_acom (((step_const ⊤)^^1) (⊥⇩c test5_const))"
value "show_acom (((step_const ⊤)^^2) (⊥⇩c test5_const))"
value "show_acom (((step_const ⊤)^^3) (⊥⇩c test5_const))"
value "show_acom (((step_const ⊤)^^4) (⊥⇩c test5_const))"
value "show_acom (((step_const ⊤)^^5) (⊥⇩c test5_const))"
value "show_acom_opt (AI_const test5_const)"
value "show_acom (((step_const ⊤)^^0) (⊥⇩c test6_const))"
value "show_acom (((step_const ⊤)^^1) (⊥⇩c test6_const))"
value "show_acom (((step_const ⊤)^^2) (⊥⇩c test6_const))"
value "show_acom (((step_const ⊤)^^3) (⊥⇩c test6_const))"
value "show_acom (((step_const ⊤)^^4) (⊥⇩c test6_const))"
value "show_acom (((step_const ⊤)^^5) (⊥⇩c test6_const))"
value "show_acom (((step_const ⊤)^^6) (⊥⇩c test6_const))"
value "show_acom (((step_const ⊤)^^7) (⊥⇩c test6_const))"
value "show_acom (((step_const ⊤)^^8) (⊥⇩c test6_const))"
value "show_acom (((step_const ⊤)^^9) (⊥⇩c test6_const))"
value "show_acom (((step_const ⊤)^^10) (⊥⇩c test6_const))"
value "show_acom (((step_const ⊤)^^11) (⊥⇩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 (standard, goal_cases)
case 1 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 "∃c'. AI_const c = Some c'"
by(rule AI_Some_measure[OF measure_const measure_const_eq])
end