| author | wenzelm | 
| Thu, 16 May 2013 19:41:41 +0200 | |
| changeset 52039 | d0ba73d11e32 | 
| parent 51826 | 054a40461449 | 
| child 52504 | 52cd8bebc3b6 | 
| permissions | -rw-r--r-- | 
| 47613 | 1  | 
(* Author: Tobias Nipkow *)  | 
2  | 
||
3  | 
theory Abs_Int1_const  | 
|
4  | 
imports Abs_Int1  | 
|
5  | 
begin  | 
|
6  | 
||
7  | 
subsection "Constant Propagation"  | 
|
8  | 
||
9  | 
datatype const = Const val | Any  | 
|
10  | 
||
11  | 
fun \<gamma>_const where  | 
|
| 51801 | 12  | 
"\<gamma>_const (Const i) = {i}" |
 | 
| 47613 | 13  | 
"\<gamma>_const (Any) = UNIV"  | 
14  | 
||
15  | 
fun plus_const where  | 
|
| 51801 | 16  | 
"plus_const (Const i) (Const j) = Const(i+j)" |  | 
| 47613 | 17  | 
"plus_const _ _ = Any"  | 
18  | 
||
19  | 
lemma plus_const_cases: "plus_const a1 a2 =  | 
|
| 51801 | 20  | 
(case (a1,a2) of (Const i, Const j) \<Rightarrow> Const(i+j) | _ \<Rightarrow> Any)"  | 
| 47613 | 21  | 
by(auto split: prod.split const.split)  | 
22  | 
||
| 51826 | 23  | 
instantiation const :: semilattice_sup_top  | 
| 47613 | 24  | 
begin  | 
25  | 
||
| 51801 | 26  | 
fun less_eq_const where "x \<le> y = (y = Any | x=y)"  | 
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
27  | 
|
| 51801 | 28  | 
definition "x < (y::const) = (x \<le> y & \<not> y \<le> x)"  | 
| 47613 | 29  | 
|
| 51802 | 30  | 
fun sup_const where "x \<squnion> y = (if x=y then x else Any)"  | 
| 47613 | 31  | 
|
32  | 
definition "\<top> = Any"  | 
|
33  | 
||
34  | 
instance  | 
|
35  | 
proof  | 
|
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
36  | 
case goal1 thus ?case by (rule less_const_def)  | 
| 
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
37  | 
next  | 
| 
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
38  | 
case goal2 show ?case by (cases x) simp_all  | 
| 47613 | 39  | 
next  | 
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
40  | 
case goal3 thus ?case by(cases z, cases y, cases x, simp_all)  | 
| 47613 | 41  | 
next  | 
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
42  | 
case goal4 thus ?case by(cases x, cases y, simp_all, cases y, simp_all)  | 
| 47613 | 43  | 
next  | 
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
44  | 
case goal6 thus ?case by(cases x, cases y, simp_all)  | 
| 
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
45  | 
next  | 
| 51389 | 46  | 
case goal5 thus ?case by(cases y, cases x, simp_all)  | 
| 47613 | 47  | 
next  | 
| 51389 | 48  | 
case goal7 thus ?case by(cases z, cases y, cases x, simp_all)  | 
| 47613 | 49  | 
next  | 
| 51389 | 50  | 
case goal8 thus ?case by(simp add: top_const_def)  | 
| 47613 | 51  | 
qed  | 
52  | 
||
53  | 
end  | 
|
54  | 
||
55  | 
||
56  | 
interpretation Val_abs  | 
|
57  | 
where \<gamma> = \<gamma>_const and num' = Const and plus' = plus_const  | 
|
58  | 
proof  | 
|
59  | 
case goal1 thus ?case  | 
|
60  | 
by(cases a, cases b, simp, simp, cases b, simp, simp)  | 
|
61  | 
next  | 
|
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
62  | 
case goal2 show ?case by(simp add: top_const_def)  | 
| 47613 | 63  | 
next  | 
64  | 
case goal3 show ?case by simp  | 
|
65  | 
next  | 
|
66  | 
case goal4 thus ?case  | 
|
67  | 
by(auto simp: plus_const_cases split: const.split)  | 
|
68  | 
qed  | 
|
69  | 
||
70  | 
interpretation Abs_Int  | 
|
71  | 
where \<gamma> = \<gamma>_const and num' = Const and plus' = plus_const  | 
|
72  | 
defines AI_const is AI and step_const is step' and aval'_const is aval'  | 
|
73  | 
..  | 
|
74  | 
||
75  | 
||
76  | 
subsubsection "Tests"  | 
|
77  | 
||
| 
51711
 
df3426139651
complete revision: finally got rid of annoying L-predicate
 
nipkow 
parents: 
51389 
diff
changeset
 | 
78  | 
definition "steps c i = (step_const \<top> ^^ i) (bot c)"  | 
| 47613 | 79  | 
|
80  | 
value "show_acom (steps test1_const 0)"  | 
|
81  | 
value "show_acom (steps test1_const 1)"  | 
|
82  | 
value "show_acom (steps test1_const 2)"  | 
|
83  | 
value "show_acom (steps test1_const 3)"  | 
|
| 50995 | 84  | 
value "show_acom (the(AI_const test1_const))"  | 
| 47613 | 85  | 
|
| 50995 | 86  | 
value "show_acom (the(AI_const test2_const))"  | 
87  | 
value "show_acom (the(AI_const test3_const))"  | 
|
| 47613 | 88  | 
|
89  | 
value "show_acom (steps test4_const 0)"  | 
|
90  | 
value "show_acom (steps test4_const 1)"  | 
|
91  | 
value "show_acom (steps test4_const 2)"  | 
|
92  | 
value "show_acom (steps test4_const 3)"  | 
|
| 49188 | 93  | 
value "show_acom (steps test4_const 4)"  | 
| 50995 | 94  | 
value "show_acom (the(AI_const test4_const))"  | 
| 47613 | 95  | 
|
96  | 
value "show_acom (steps test5_const 0)"  | 
|
97  | 
value "show_acom (steps test5_const 1)"  | 
|
98  | 
value "show_acom (steps test5_const 2)"  | 
|
99  | 
value "show_acom (steps test5_const 3)"  | 
|
100  | 
value "show_acom (steps test5_const 4)"  | 
|
101  | 
value "show_acom (steps test5_const 5)"  | 
|
| 49188 | 102  | 
value "show_acom (steps test5_const 6)"  | 
| 50995 | 103  | 
value "show_acom (the(AI_const test5_const))"  | 
| 47613 | 104  | 
|
105  | 
value "show_acom (steps test6_const 0)"  | 
|
106  | 
value "show_acom (steps test6_const 1)"  | 
|
107  | 
value "show_acom (steps test6_const 2)"  | 
|
108  | 
value "show_acom (steps test6_const 3)"  | 
|
109  | 
value "show_acom (steps test6_const 4)"  | 
|
110  | 
value "show_acom (steps test6_const 5)"  | 
|
111  | 
value "show_acom (steps test6_const 6)"  | 
|
112  | 
value "show_acom (steps test6_const 7)"  | 
|
113  | 
value "show_acom (steps test6_const 8)"  | 
|
114  | 
value "show_acom (steps test6_const 9)"  | 
|
115  | 
value "show_acom (steps test6_const 10)"  | 
|
116  | 
value "show_acom (steps test6_const 11)"  | 
|
| 49188 | 117  | 
value "show_acom (steps test6_const 12)"  | 
118  | 
value "show_acom (steps test6_const 13)"  | 
|
| 50995 | 119  | 
value "show_acom (the(AI_const test6_const))"  | 
| 47613 | 120  | 
|
121  | 
||
122  | 
text{* Monotonicity: *}
 | 
|
123  | 
||
124  | 
interpretation Abs_Int_mono  | 
|
125  | 
where \<gamma> = \<gamma>_const and num' = Const and plus' = plus_const  | 
|
126  | 
proof  | 
|
127  | 
case goal1 thus ?case  | 
|
128  | 
by(auto simp: plus_const_cases split: const.split)  | 
|
129  | 
qed  | 
|
130  | 
||
131  | 
text{* Termination: *}
 | 
|
132  | 
||
| 51803 | 133  | 
definition m_const :: "const \<Rightarrow> nat" where  | 
134  | 
"m_const x = (if x = Any then 0 else 1)"  | 
|
| 47613 | 135  | 
|
136  | 
interpretation Abs_Int_measure  | 
|
137  | 
where \<gamma> = \<gamma>_const and num' = Const and plus' = plus_const  | 
|
| 49433 | 138  | 
and m = m_const and h = "1"  | 
| 47613 | 139  | 
proof  | 
140  | 
case goal1 thus ?case by(auto simp: m_const_def split: const.splits)  | 
|
141  | 
next  | 
|
| 51372 | 142  | 
case goal2 thus ?case by(auto simp: m_const_def less_const_def split: const.splits)  | 
| 47613 | 143  | 
qed  | 
144  | 
||
145  | 
thm AI_Some_measure  | 
|
146  | 
||
147  | 
end  |