| author | blanchet | 
| Tue, 09 Sep 2014 20:51:36 +0200 | |
| changeset 58255 | 9dfe8506c04d | 
| parent 58249 | 180f1b3508ed | 
| child 58310 | 91ea607a34d8 | 
| 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  | 
||
| 
58249
 
180f1b3508ed
use 'datatype_new' (soon to be renamed 'datatype') in Isabelle's libraries
 
blanchet 
parents: 
56927 
diff
changeset
 | 
9  | 
datatype_new parity = Even | Odd | Either  | 
| 47613 | 10  | 
|
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
11  | 
text{* Instantiation of class @{class order} with type @{typ parity}: *}
 | 
| 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  | 
||
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
16  | 
text{* First the definition of the interface function @{text"\<le>"}. Note that
 | 
| 
 
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
 | 
| 
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
19  | 
less_eq_parity_def}. Inside the definition the symbolic names can be used. *}  | 
| 
 
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  | 
|
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
24  | 
text{* We also need @{text"<"}, which is defined canonically: *}
 | 
| 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  | 
|
| 51625 | 29  | 
text{*\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  | 
33  | 
necessary proof obligations. *}  | 
|
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  | 
||
| 51826 | 50  | 
text{* Instantiation of class @{class semilattice_sup_top} with type @{typ parity}: *}
 | 
| 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  | 
||
61  | 
text{* Now the instance proof. This time we take a lazy shortcut: we do not
 | 
|
62  | 
write out the proof obligations but use the @{text goali} primitive to refer
 | 
|
63  | 
to the assumptions of subgoal i and @{text "case?"} to refer to the
 | 
|
64  | 
conclusion of subgoal i. The class axioms are presented in the same order as  | 
|
| 51625 | 65  | 
in the class definition. Warning: this is brittle! *}  | 
| 47613 | 66  | 
|
67  | 
instance  | 
|
68  | 
proof  | 
|
| 51389 | 69  | 
case goal1 (*sup1*) show ?case by(auto simp: less_eq_parity_def sup_parity_def)  | 
| 47613 | 70  | 
next  | 
| 51389 | 71  | 
case goal2 (*sup2*) show ?case by(auto simp: less_eq_parity_def sup_parity_def)  | 
| 47613 | 72  | 
next  | 
| 51389 | 73  | 
case goal3 (*sup least*) thus ?case by(auto simp: less_eq_parity_def sup_parity_def)  | 
| 47613 | 74  | 
next  | 
| 51389 | 75  | 
case goal4 (*top*) show ?case by(auto simp: less_eq_parity_def top_parity_def)  | 
| 47613 | 76  | 
qed  | 
77  | 
||
78  | 
end  | 
|
79  | 
||
80  | 
||
81  | 
text{* Now we define the functions used for instantiating the abstract
 | 
|
82  | 
interpretation locales. Note that the Isabelle terminology is  | 
|
83  | 
\emph{interpretation}, not \emph{instantiation} of locales, but we use
 | 
|
84  | 
instantiation to avoid confusion with abstract interpretation. *}  | 
|
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  | 
||
102  | 
text{* First we instantiate the abstract value interface and prove that the
 | 
|
103  | 
functions on type @{typ parity} have all the necessary properties: *}
 | 
|
104  | 
||
| 
55599
 
6535c537b243
aggiornamento for "interpretation with definitions": operate uniformly on theory and locale level under the brand of "permanent interpretation"
 
haftmann 
parents: 
52525 
diff
changeset
 | 
105  | 
permanent_interpretation Val_semilattice  | 
| 47613 | 106  | 
where \<gamma> = \<gamma>_parity and num' = num_parity and plus' = plus_parity  | 
107  | 
proof txt{* of the locale axioms *}
 | 
|
108  | 
fix a b :: parity  | 
|
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
109  | 
assume "a \<le> b" thus "\<gamma>_parity a \<subseteq> \<gamma>_parity b"  | 
| 
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
110  | 
by(auto simp: less_eq_parity_def)  | 
| 47613 | 111  | 
next txt{* The rest in the lazy, implicit way *}
 | 
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
112  | 
case goal2 show ?case by(auto simp: top_parity_def)  | 
| 47613 | 113  | 
next  | 
114  | 
case goal3 show ?case by auto  | 
|
115  | 
next  | 
|
116  | 
  txt{* Warning: this subproof refers to the names @{text a1} and @{text a2}
 | 
|
117  | 
from the statement of the axiom. *}  | 
|
118  | 
case goal4 thus ?case  | 
|
| 52525 | 119  | 
by (induction a1 a2 rule: plus_parity.induct) (auto simp add:mod_add_eq)  | 
| 47613 | 120  | 
qed  | 
121  | 
||
122  | 
text{* Instantiating the abstract interpretation locale requires no more
 | 
|
123  | 
proofs (they happened in the instatiation above) but delivers the  | 
|
| 49344 | 124  | 
instantiated abstract interpreter which we call @{text AI_parity}: *}
 | 
| 47613 | 125  | 
|
| 
55599
 
6535c537b243
aggiornamento for "interpretation with definitions": operate uniformly on theory and locale level under the brand of "permanent interpretation"
 
haftmann 
parents: 
52525 
diff
changeset
 | 
126  | 
permanent_interpretation Abs_Int  | 
| 47613 | 127  | 
where \<gamma> = \<gamma>_parity and num' = num_parity and plus' = plus_parity  | 
| 
55600
 
3c7610b8dcfc
more convenient syntax for permanent interpretation
 
haftmann 
parents: 
55599 
diff
changeset
 | 
128  | 
defining aval_parity = aval' and step_parity = step' and AI_parity = AI  | 
| 47613 | 129  | 
..  | 
130  | 
||
131  | 
||
132  | 
subsubsection "Tests"  | 
|
133  | 
||
134  | 
definition "test1_parity =  | 
|
| 
52046
 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 
nipkow 
parents: 
51826 
diff
changeset
 | 
135  | 
''x'' ::= N 1;;  | 
| 47613 | 136  | 
WHILE Less (V ''x'') (N 100) DO ''x'' ::= Plus (V ''x'') (N 2)"  | 
| 56927 | 137  | 
value "show_acom (the(AI_parity test1_parity))"  | 
| 47613 | 138  | 
|
139  | 
definition "test2_parity =  | 
|
| 
52046
 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 
nipkow 
parents: 
51826 
diff
changeset
 | 
140  | 
''x'' ::= N 1;;  | 
| 47613 | 141  | 
WHILE Less (V ''x'') (N 100) DO ''x'' ::= Plus (V ''x'') (N 3)"  | 
142  | 
||
| 
51711
 
df3426139651
complete revision: finally got rid of annoying L-predicate
 
nipkow 
parents: 
51625 
diff
changeset
 | 
143  | 
definition "steps c i = ((step_parity \<top>) ^^ i) (bot c)"  | 
| 47613 | 144  | 
|
145  | 
value "show_acom (steps test2_parity 0)"  | 
|
146  | 
value "show_acom (steps test2_parity 1)"  | 
|
147  | 
value "show_acom (steps test2_parity 2)"  | 
|
148  | 
value "show_acom (steps test2_parity 3)"  | 
|
149  | 
value "show_acom (steps test2_parity 4)"  | 
|
150  | 
value "show_acom (steps test2_parity 5)"  | 
|
| 49188 | 151  | 
value "show_acom (steps test2_parity 6)"  | 
| 50995 | 152  | 
value "show_acom (the(AI_parity test2_parity))"  | 
| 47613 | 153  | 
|
154  | 
||
155  | 
subsubsection "Termination"  | 
|
156  | 
||
| 
55599
 
6535c537b243
aggiornamento for "interpretation with definitions": operate uniformly on theory and locale level under the brand of "permanent interpretation"
 
haftmann 
parents: 
52525 
diff
changeset
 | 
157  | 
permanent_interpretation Abs_Int_mono  | 
| 47613 | 158  | 
where \<gamma> = \<gamma>_parity and num' = num_parity and plus' = plus_parity  | 
159  | 
proof  | 
|
160  | 
case goal1 thus ?case  | 
|
| 52525 | 161  | 
by(induction a1 a2 rule: plus_parity.induct)  | 
162  | 
(auto simp add:less_eq_parity_def)  | 
|
| 47613 | 163  | 
qed  | 
164  | 
||
165  | 
definition m_parity :: "parity \<Rightarrow> nat" where  | 
|
| 51389 | 166  | 
"m_parity x = (if x = Either then 0 else 1)"  | 
| 47613 | 167  | 
|
| 
55599
 
6535c537b243
aggiornamento for "interpretation with definitions": operate uniformly on theory and locale level under the brand of "permanent interpretation"
 
haftmann 
parents: 
52525 
diff
changeset
 | 
168  | 
permanent_interpretation Abs_Int_measure  | 
| 47613 | 169  | 
where \<gamma> = \<gamma>_parity and num' = num_parity and plus' = plus_parity  | 
| 49433 | 170  | 
and m = m_parity and h = "1"  | 
| 47613 | 171  | 
proof  | 
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
172  | 
case goal1 thus ?case by(auto simp add: m_parity_def less_eq_parity_def)  | 
| 47613 | 173  | 
next  | 
| 51372 | 174  | 
case goal2 thus ?case by(auto simp add: m_parity_def less_eq_parity_def less_parity_def)  | 
| 47613 | 175  | 
qed  | 
176  | 
||
177  | 
thm AI_Some_measure  | 
|
178  | 
||
179  | 
end  |