author | wenzelm |
Fri, 27 Jul 2018 22:23:37 +0200 | |
changeset 68695 | 9072bfd24d8f |
parent 67406 | 23307fd33906 |
child 68778 | 4566bac4517d |
permissions | -rw-r--r-- |
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
1 |
(* Author: Tobias Nipkow *) |
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
2 |
|
47613 | 3 |
theory Abs_Int1_parity |
4 |
imports Abs_Int1 |
|
5 |
begin |
|
6 |
||
7 |
subsection "Parity Analysis" |
|
8 |
||
58310 | 9 |
datatype parity = Even | Odd | Either |
47613 | 10 |
|
67406 | 11 |
text\<open>Instantiation of class @{class order} with type @{typ parity}:\<close> |
47613 | 12 |
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
13 |
instantiation parity :: order |
47613 | 14 |
begin |
15 |
||
67406 | 16 |
text\<open>First the definition of the interface function @{text"\<le>"}. Note that |
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
17 |
the header of the definition must refer to the ascii name @{const less_eq} of the |
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
18 |
constants as @{text less_eq_parity} and the definition is named @{text |
67406 | 19 |
less_eq_parity_def}. Inside the definition the symbolic names can be used.\<close> |
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
20 |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
21 |
definition less_eq_parity where |
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
22 |
"x \<le> y = (y = Either \<or> x=y)" |
47613 | 23 |
|
67406 | 24 |
text\<open>We also need @{text"<"}, which is defined canonically:\<close> |
47613 | 25 |
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
26 |
definition less_parity where |
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
27 |
"x < y = (x \<le> y \<and> \<not> y \<le> (x::parity))" |
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
28 |
|
67406 | 29 |
text\<open>\noindent(The type annotation is necessary to fix the type of the polymorphic predicates.) |
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
30 |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
31 |
Now the instance proof, i.e.\ the proof that the definition fulfills |
47613 | 32 |
the axioms (assumptions) of the class. The initial proof-step generates the |
67406 | 33 |
necessary proof obligations.\<close> |
47613 | 34 |
|
35 |
instance |
|
36 |
proof |
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
37 |
fix x::parity show "x \<le> x" by(auto simp: less_eq_parity_def) |
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
38 |
next |
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
39 |
fix x y z :: parity assume "x \<le> y" "y \<le> z" thus "x \<le> z" |
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
40 |
by(auto simp: less_eq_parity_def) |
47613 | 41 |
next |
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
42 |
fix x y :: parity assume "x \<le> y" "y \<le> x" thus "x = y" |
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
43 |
by(auto simp: less_eq_parity_def) |
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
44 |
next |
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
45 |
fix x y :: parity show "(x < y) = (x \<le> y \<and> \<not> y \<le> x)" by(rule less_parity_def) |
47613 | 46 |
qed |
47 |
||
48 |
end |
|
49 |
||
67406 | 50 |
text\<open>Instantiation of class @{class semilattice_sup_top} with type @{typ parity}:\<close> |
47613 | 51 |
|
51826 | 52 |
instantiation parity :: semilattice_sup_top |
47613 | 53 |
begin |
54 |
||
51389 | 55 |
definition sup_parity where |
51624 | 56 |
"x \<squnion> y = (if x = y then x else Either)" |
47613 | 57 |
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
58 |
definition top_parity where |
47613 | 59 |
"\<top> = Either" |
60 |
||
67406 | 61 |
text\<open>Now the instance proof. This time we take a shortcut with the help of |
61179 | 62 |
proof method @{text goal_cases}: it creates cases 1 ... n for the subgoals |
63 |
1 ... n; in case i, i is also the name of the assumptions of subgoal i and |
|
64 |
@{text "case?"} refers to the conclusion of subgoal i. |
|
67406 | 65 |
The class axioms are presented in the same order as in the class definition.\<close> |
47613 | 66 |
|
67 |
instance |
|
61179 | 68 |
proof (standard, goal_cases) |
69 |
case 1 (*sup1*) show ?case by(auto simp: less_eq_parity_def sup_parity_def) |
|
47613 | 70 |
next |
61179 | 71 |
case 2 (*sup2*) show ?case by(auto simp: less_eq_parity_def sup_parity_def) |
47613 | 72 |
next |
61179 | 73 |
case 3 (*sup least*) thus ?case by(auto simp: less_eq_parity_def sup_parity_def) |
47613 | 74 |
next |
61179 | 75 |
case 4 (*top*) show ?case by(auto simp: less_eq_parity_def top_parity_def) |
47613 | 76 |
qed |
77 |
||
78 |
end |
|
79 |
||
80 |
||
67406 | 81 |
text\<open>Now we define the functions used for instantiating the abstract |
47613 | 82 |
interpretation locales. Note that the Isabelle terminology is |
83 |
\emph{interpretation}, not \emph{instantiation} of locales, but we use |
|
67406 | 84 |
instantiation to avoid confusion with abstract interpretation.\<close> |
47613 | 85 |
|
86 |
fun \<gamma>_parity :: "parity \<Rightarrow> val set" where |
|
87 |
"\<gamma>_parity Even = {i. i mod 2 = 0}" | |
|
88 |
"\<gamma>_parity Odd = {i. i mod 2 = 1}" | |
|
89 |
"\<gamma>_parity Either = UNIV" |
|
90 |
||
91 |
fun num_parity :: "val \<Rightarrow> parity" where |
|
92 |
"num_parity i = (if i mod 2 = 0 then Even else Odd)" |
|
93 |
||
94 |
fun plus_parity :: "parity \<Rightarrow> parity \<Rightarrow> parity" where |
|
95 |
"plus_parity Even Even = Even" | |
|
96 |
"plus_parity Odd Odd = Even" | |
|
97 |
"plus_parity Even Odd = Odd" | |
|
98 |
"plus_parity Odd Even = Odd" | |
|
99 |
"plus_parity Either y = Either" | |
|
100 |
"plus_parity x Either = Either" |
|
101 |
||
67406 | 102 |
text\<open>First we instantiate the abstract value interface and prove that the |
103 |
functions on type @{typ parity} have all the necessary properties:\<close> |
|
47613 | 104 |
|
61890
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
haftmann
parents:
61671
diff
changeset
|
105 |
global_interpretation Val_semilattice |
47613 | 106 |
where \<gamma> = \<gamma>_parity and num' = num_parity and plus' = plus_parity |
67406 | 107 |
proof (standard, goal_cases) txt\<open>subgoals are the locale axioms\<close> |
61179 | 108 |
case 1 thus ?case by(auto simp: less_eq_parity_def) |
47613 | 109 |
next |
61179 | 110 |
case 2 show ?case by(auto simp: top_parity_def) |
47613 | 111 |
next |
61179 | 112 |
case 3 show ?case by auto |
113 |
next |
|
114 |
case (4 _ a1 _ a2) thus ?case |
|
64593
50c715579715
reoriented congruence rules in non-explosive direction
haftmann
parents:
61890
diff
changeset
|
115 |
by (induction a1 a2 rule: plus_parity.induct) |
50c715579715
reoriented congruence rules in non-explosive direction
haftmann
parents:
61890
diff
changeset
|
116 |
(auto simp add: mod_add_eq [symmetric]) |
47613 | 117 |
qed |
118 |
||
67406 | 119 |
text\<open>In case 4 we needed to refer to particular variables. |
61179 | 120 |
Writing (i x y z) fixes the names of the variables in case i to be x, y and z |
121 |
in the left-to-right order in which the variables occur in the subgoal. |
|
67406 | 122 |
Underscores are anonymous placeholders for variable names we don't care to fix.\<close> |
61179 | 123 |
|
67406 | 124 |
text\<open>Instantiating the abstract interpretation locale requires no more |
47613 | 125 |
proofs (they happened in the instatiation above) but delivers the |
67406 | 126 |
instantiated abstract interpreter which we call @{text AI_parity}:\<close> |
47613 | 127 |
|
61890
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
haftmann
parents:
61671
diff
changeset
|
128 |
global_interpretation Abs_Int |
47613 | 129 |
where \<gamma> = \<gamma>_parity and num' = num_parity and plus' = plus_parity |
61671
20d4cd2ceab2
prefer "rewrites" and "defines" to note rewrite morphisms
haftmann
parents:
61179
diff
changeset
|
130 |
defines aval_parity = aval' and step_parity = step' and AI_parity = AI |
47613 | 131 |
.. |
132 |
||
133 |
||
134 |
subsubsection "Tests" |
|
135 |
||
136 |
definition "test1_parity = |
|
52046
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents:
51826
diff
changeset
|
137 |
''x'' ::= N 1;; |
47613 | 138 |
WHILE Less (V ''x'') (N 100) DO ''x'' ::= Plus (V ''x'') (N 2)" |
56927 | 139 |
value "show_acom (the(AI_parity test1_parity))" |
47613 | 140 |
|
141 |
definition "test2_parity = |
|
52046
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents:
51826
diff
changeset
|
142 |
''x'' ::= N 1;; |
47613 | 143 |
WHILE Less (V ''x'') (N 100) DO ''x'' ::= Plus (V ''x'') (N 3)" |
144 |
||
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51625
diff
changeset
|
145 |
definition "steps c i = ((step_parity \<top>) ^^ i) (bot c)" |
47613 | 146 |
|
147 |
value "show_acom (steps test2_parity 0)" |
|
148 |
value "show_acom (steps test2_parity 1)" |
|
149 |
value "show_acom (steps test2_parity 2)" |
|
150 |
value "show_acom (steps test2_parity 3)" |
|
151 |
value "show_acom (steps test2_parity 4)" |
|
152 |
value "show_acom (steps test2_parity 5)" |
|
49188 | 153 |
value "show_acom (steps test2_parity 6)" |
50995 | 154 |
value "show_acom (the(AI_parity test2_parity))" |
47613 | 155 |
|
156 |
||
157 |
subsubsection "Termination" |
|
158 |
||
61890
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
haftmann
parents:
61671
diff
changeset
|
159 |
global_interpretation Abs_Int_mono |
47613 | 160 |
where \<gamma> = \<gamma>_parity and num' = num_parity and plus' = plus_parity |
61179 | 161 |
proof (standard, goal_cases) |
162 |
case (1 _ a1 _ a2) thus ?case |
|
52525 | 163 |
by(induction a1 a2 rule: plus_parity.induct) |
164 |
(auto simp add:less_eq_parity_def) |
|
47613 | 165 |
qed |
166 |
||
167 |
definition m_parity :: "parity \<Rightarrow> nat" where |
|
51389 | 168 |
"m_parity x = (if x = Either then 0 else 1)" |
47613 | 169 |
|
61890
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
haftmann
parents:
61671
diff
changeset
|
170 |
global_interpretation Abs_Int_measure |
47613 | 171 |
where \<gamma> = \<gamma>_parity and num' = num_parity and plus' = plus_parity |
49433 | 172 |
and m = m_parity and h = "1" |
61179 | 173 |
proof (standard, goal_cases) |
174 |
case 1 thus ?case by(auto simp add: m_parity_def less_eq_parity_def) |
|
47613 | 175 |
next |
61179 | 176 |
case 2 thus ?case by(auto simp add: m_parity_def less_eq_parity_def less_parity_def) |
47613 | 177 |
qed |
178 |
||
179 |
thm AI_Some_measure |
|
180 |
||
181 |
end |