author | haftmann |
Sat, 24 Dec 2011 15:53:10 +0100 | |
changeset 45970 | b6d0cff57d96 |
parent 45694 | 4a8743618257 |
child 46104 | eb85282db54e |
permissions | -rw-r--r-- |
33197 | 1 |
(* Title: HOL/Nitpick_Examples/Manual_Nits.thy |
2 |
Author: Jasmin Blanchette, TU Muenchen |
|
45035 | 3 |
Copyright 2009-2011 |
33197 | 4 |
|
5 |
Examples from the Nitpick manual. |
|
6 |
*) |
|
7 |
||
8 |
header {* Examples from the Nitpick Manual *} |
|
9 |
||
37477
e482320bcbfe
adjusted Nitpick examples to latest changes + make them slightly faster
blanchet
parents:
36268
diff
changeset
|
10 |
(* The "expect" arguments to Nitpick in this theory and the other example |
e482320bcbfe
adjusted Nitpick examples to latest changes + make them slightly faster
blanchet
parents:
36268
diff
changeset
|
11 |
theories are there so that the example can also serve as a regression test |
e482320bcbfe
adjusted Nitpick examples to latest changes + make them slightly faster
blanchet
parents:
36268
diff
changeset
|
12 |
suite. *) |
e482320bcbfe
adjusted Nitpick examples to latest changes + make them slightly faster
blanchet
parents:
36268
diff
changeset
|
13 |
|
33197 | 14 |
theory Manual_Nits |
41413
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
wenzelm
parents:
41278
diff
changeset
|
15 |
imports Main "~~/src/HOL/Library/Quotient_Product" RealDef |
33197 | 16 |
begin |
17 |
||
45053 | 18 |
chapter {* 2. First Steps *} |
33197 | 19 |
|
41278
8e1cde88aae6
added a timestamp to Nitpick in verbose mode for debugging purposes;
blanchet
parents:
40341
diff
changeset
|
20 |
nitpick_params [verbose, sat_solver = MiniSat_JNI, max_threads = 1, |
42208
02513eb26eb7
raised timeouts further, for SML/NJ -- because of variations in machines/compilers, fixed timeouts can merely prevent non-termination, not enforce particular performance characteristics.
krauss
parents:
41413
diff
changeset
|
21 |
timeout = 240] |
33197 | 22 |
|
45053 | 23 |
subsection {* 2.1. Propositional Logic *} |
33197 | 24 |
|
25 |
lemma "P \<longleftrightarrow> Q" |
|
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
26 |
nitpick [expect = genuine] |
33197 | 27 |
apply auto |
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
28 |
nitpick [expect = genuine] 1 |
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
29 |
nitpick [expect = genuine] 2 |
33197 | 30 |
oops |
31 |
||
45053 | 32 |
subsection {* 2.2. Type Variables *} |
33197 | 33 |
|
34 |
lemma "P x \<Longrightarrow> P (THE y. P y)" |
|
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
35 |
nitpick [verbose, expect = genuine] |
33197 | 36 |
oops |
37 |
||
45053 | 38 |
subsection {* 2.3. Constants *} |
33197 | 39 |
|
40 |
lemma "P x \<Longrightarrow> P (THE y. P y)" |
|
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
41 |
nitpick [show_consts, expect = genuine] |
39362
ee65900bfced
adapt examples to latest Nitpick changes + speed them up a little bit
blanchet
parents:
39302
diff
changeset
|
42 |
nitpick [dont_specialize, show_consts, expect = genuine] |
33197 | 43 |
oops |
44 |
||
45 |
lemma "\<exists>!x. P x \<Longrightarrow> P (THE y. P y)" |
|
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
46 |
nitpick [expect = none] |
42959 | 47 |
nitpick [card 'a = 1\<emdash>50, expect = none] |
33197 | 48 |
(* sledgehammer *) |
49 |
apply (metis the_equality) |
|
50 |
done |
|
51 |
||
45053 | 52 |
subsection {* 2.4. Skolemization *} |
33197 | 53 |
|
54 |
lemma "\<exists>g. \<forall>x. g (f x) = x \<Longrightarrow> \<forall>y. \<exists>x. y = f x" |
|
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
55 |
nitpick [expect = genuine] |
33197 | 56 |
oops |
57 |
||
58 |
lemma "\<exists>x. \<forall>f. f x = x" |
|
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
59 |
nitpick [expect = genuine] |
33197 | 60 |
oops |
61 |
||
62 |
lemma "refl r \<Longrightarrow> sym r" |
|
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
63 |
nitpick [expect = genuine] |
33197 | 64 |
oops |
65 |
||
45053 | 66 |
subsection {* 2.5. Natural Numbers and Integers *} |
33197 | 67 |
|
68 |
lemma "\<lbrakk>i \<le> j; n \<le> (m\<Colon>int)\<rbrakk> \<Longrightarrow> i * n + j * m \<le> i * m + j * n" |
|
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
69 |
nitpick [expect = genuine] |
33197 | 70 |
oops |
71 |
||
72 |
lemma "\<forall>n. Suc n \<noteq> n \<Longrightarrow> P" |
|
42421
6bc725d60593
increase "auto"'s timeout in example to help SML/NJ
blanchet
parents:
42208
diff
changeset
|
73 |
nitpick [card nat = 100, check_potential, tac_timeout = 5, expect = genuine] |
33197 | 74 |
oops |
75 |
||
76 |
lemma "P Suc" |
|
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
77 |
nitpick [expect = none] |
33197 | 78 |
oops |
79 |
||
80 |
lemma "P (op +\<Colon>nat\<Rightarrow>nat\<Rightarrow>nat)" |
|
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
81 |
nitpick [card nat = 1, expect = genuine] |
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
82 |
nitpick [card nat = 2, expect = none] |
33197 | 83 |
oops |
84 |
||
45053 | 85 |
subsection {* 2.6. Inductive Datatypes *} |
33197 | 86 |
|
87 |
lemma "hd (xs @ [y, y]) = hd xs" |
|
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
88 |
nitpick [expect = genuine] |
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
89 |
nitpick [show_consts, show_datatypes, expect = genuine] |
33197 | 90 |
oops |
91 |
||
92 |
lemma "\<lbrakk>length xs = 1; length ys = 1\<rbrakk> \<Longrightarrow> xs = ys" |
|
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
93 |
nitpick [show_datatypes, expect = genuine] |
33197 | 94 |
oops |
95 |
||
45053 | 96 |
subsection {* 2.7. Typedefs, Records, Rationals, and Reals *} |
33197 | 97 |
|
45694
4a8743618257
prefer typedef without extra definition and alternative name;
wenzelm
parents:
45053
diff
changeset
|
98 |
definition "three = {0\<Colon>nat, 1, 2}" |
4a8743618257
prefer typedef without extra definition and alternative name;
wenzelm
parents:
45053
diff
changeset
|
99 |
typedef (open) three = three |
4a8743618257
prefer typedef without extra definition and alternative name;
wenzelm
parents:
45053
diff
changeset
|
100 |
unfolding three_def by blast |
33197 | 101 |
|
102 |
definition A :: three where "A \<equiv> Abs_three 0" |
|
103 |
definition B :: three where "B \<equiv> Abs_three 1" |
|
104 |
definition C :: three where "C \<equiv> Abs_three 2" |
|
105 |
||
106 |
lemma "\<lbrakk>P A; P B\<rbrakk> \<Longrightarrow> P x" |
|
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
107 |
nitpick [show_datatypes, expect = genuine] |
33197 | 108 |
oops |
109 |
||
35284
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents:
35185
diff
changeset
|
110 |
fun my_int_rel where |
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents:
35185
diff
changeset
|
111 |
"my_int_rel (x, y) (u, v) = (x + v = u + y)" |
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents:
35185
diff
changeset
|
112 |
|
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents:
35185
diff
changeset
|
113 |
quotient_type my_int = "nat \<times> nat" / my_int_rel |
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
114 |
by (auto simp add: equivp_def fun_eq_iff) |
35284
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents:
35185
diff
changeset
|
115 |
|
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents:
35185
diff
changeset
|
116 |
definition add_raw where |
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents:
35185
diff
changeset
|
117 |
"add_raw \<equiv> \<lambda>(x, y) (u, v). (x + (u\<Colon>nat), y + (v\<Colon>nat))" |
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents:
35185
diff
changeset
|
118 |
|
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents:
35185
diff
changeset
|
119 |
quotient_definition "add\<Colon>my_int \<Rightarrow> my_int \<Rightarrow> my_int" is add_raw |
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents:
35185
diff
changeset
|
120 |
|
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents:
35185
diff
changeset
|
121 |
lemma "add x y = add x x" |
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
122 |
nitpick [show_datatypes, expect = genuine] |
35284
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents:
35185
diff
changeset
|
123 |
oops |
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents:
35185
diff
changeset
|
124 |
|
35711
548d3f16404b
added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents:
35710
diff
changeset
|
125 |
ML {* |
35712
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
126 |
fun my_int_postproc _ _ _ T (Const _ $ (Const _ $ t1 $ t2)) = |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
127 |
HOLogic.mk_number T (snd (HOLogic.dest_number t1) |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
128 |
- snd (HOLogic.dest_number t2)) |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
129 |
| my_int_postproc _ _ _ _ t = t |
35711
548d3f16404b
added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents:
35710
diff
changeset
|
130 |
*} |
548d3f16404b
added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents:
35710
diff
changeset
|
131 |
|
38288 | 132 |
declaration {* |
38284
9f98107ad8b4
use "declaration" instead of "setup" to register Nitpick extensions
blanchet
parents:
38242
diff
changeset
|
133 |
Nitpick_Model.register_term_postprocessor @{typ my_int} my_int_postproc |
38242 | 134 |
*} |
35711
548d3f16404b
added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents:
35710
diff
changeset
|
135 |
|
548d3f16404b
added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents:
35710
diff
changeset
|
136 |
lemma "add x y = add x x" |
548d3f16404b
added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents:
35710
diff
changeset
|
137 |
nitpick [show_datatypes] |
548d3f16404b
added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents:
35710
diff
changeset
|
138 |
oops |
548d3f16404b
added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents:
35710
diff
changeset
|
139 |
|
33197 | 140 |
record point = |
141 |
Xcoord :: int |
|
142 |
Ycoord :: int |
|
143 |
||
144 |
lemma "Xcoord (p\<Colon>point) = Xcoord (q\<Colon>point)" |
|
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
145 |
nitpick [show_datatypes, expect = genuine] |
33197 | 146 |
oops |
147 |
||
148 |
lemma "4 * x + 3 * (y\<Colon>real) \<noteq> 1 / 2" |
|
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
149 |
nitpick [show_datatypes, expect = genuine] |
33197 | 150 |
oops |
151 |
||
45053 | 152 |
subsection {* 2.8. Inductive and Coinductive Predicates *} |
33197 | 153 |
|
154 |
inductive even where |
|
155 |
"even 0" | |
|
156 |
"even n \<Longrightarrow> even (Suc (Suc n))" |
|
157 |
||
158 |
lemma "\<exists>n. even n \<and> even (Suc n)" |
|
35710 | 159 |
nitpick [card nat = 50, unary_ints, verbose, expect = potential] |
33197 | 160 |
oops |
161 |
||
35710 | 162 |
lemma "\<exists>n \<le> 49. even n \<and> even (Suc n)" |
38184 | 163 |
nitpick [card nat = 50, unary_ints, expect = genuine] |
33197 | 164 |
oops |
165 |
||
166 |
inductive even' where |
|
167 |
"even' (0\<Colon>nat)" | |
|
168 |
"even' 2" | |
|
169 |
"\<lbrakk>even' m; even' n\<rbrakk> \<Longrightarrow> even' (m + n)" |
|
170 |
||
171 |
lemma "\<exists>n \<in> {0, 2, 4, 6, 8}. \<not> even' n" |
|
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
172 |
nitpick [card nat = 10, unary_ints, verbose, show_consts, expect = genuine] |
33197 | 173 |
oops |
174 |
||
175 |
lemma "even' (n - 2) \<Longrightarrow> even' n" |
|
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
176 |
nitpick [card nat = 10, show_consts, expect = genuine] |
33197 | 177 |
oops |
178 |
||
179 |
coinductive nats where |
|
180 |
"nats (x\<Colon>nat) \<Longrightarrow> nats x" |
|
181 |
||
45970
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
45694
diff
changeset
|
182 |
lemma "nats = (\<lambda>n. n \<in> {0, 1, 2, 3, 4})" |
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
183 |
nitpick [card nat = 10, show_consts, expect = genuine] |
33197 | 184 |
oops |
185 |
||
186 |
inductive odd where |
|
187 |
"odd 1" | |
|
188 |
"\<lbrakk>odd m; even n\<rbrakk> \<Longrightarrow> odd (m + n)" |
|
189 |
||
190 |
lemma "odd n \<Longrightarrow> odd (n - 2)" |
|
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
191 |
nitpick [card nat = 10, show_consts, expect = genuine] |
33197 | 192 |
oops |
193 |
||
45053 | 194 |
subsection {* 2.9. Coinductive Datatypes *} |
33197 | 195 |
|
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35312
diff
changeset
|
196 |
(* Lazy lists are defined in Andreas Lochbihler's "Coinductive" AFP entry. Since |
38184 | 197 |
we cannot rely on its presence, we expediently provide our own |
198 |
axiomatization. The examples also work unchanged with Lochbihler's |
|
199 |
"Coinductive_List" theory. *) |
|
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35312
diff
changeset
|
200 |
|
45694
4a8743618257
prefer typedef without extra definition and alternative name;
wenzelm
parents:
45053
diff
changeset
|
201 |
definition "llist = (UNIV\<Colon>('a list + (nat \<Rightarrow> 'a)) set)" |
4a8743618257
prefer typedef without extra definition and alternative name;
wenzelm
parents:
45053
diff
changeset
|
202 |
|
4a8743618257
prefer typedef without extra definition and alternative name;
wenzelm
parents:
45053
diff
changeset
|
203 |
typedef (open) 'a llist = "llist\<Colon>('a list + (nat \<Rightarrow> 'a)) set" |
4a8743618257
prefer typedef without extra definition and alternative name;
wenzelm
parents:
45053
diff
changeset
|
204 |
unfolding llist_def by auto |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35312
diff
changeset
|
205 |
|
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
206 |
definition LNil where |
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
207 |
"LNil = Abs_llist (Inl [])" |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35312
diff
changeset
|
208 |
definition LCons where |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35312
diff
changeset
|
209 |
"LCons y ys = Abs_llist (case Rep_llist ys of |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35312
diff
changeset
|
210 |
Inl ys' \<Rightarrow> Inl (y # ys') |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35312
diff
changeset
|
211 |
| Inr f \<Rightarrow> Inr (\<lambda>n. case n of 0 \<Rightarrow> y | Suc m \<Rightarrow> f m))" |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35312
diff
changeset
|
212 |
|
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35312
diff
changeset
|
213 |
axiomatization iterates :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a llist" |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35312
diff
changeset
|
214 |
|
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35312
diff
changeset
|
215 |
lemma iterates_def [nitpick_simp]: |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35312
diff
changeset
|
216 |
"iterates f a = LCons a (iterates f (f a))" |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35312
diff
changeset
|
217 |
sorry |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35312
diff
changeset
|
218 |
|
38288 | 219 |
declaration {* |
38284
9f98107ad8b4
use "declaration" instead of "setup" to register Nitpick extensions
blanchet
parents:
38242
diff
changeset
|
220 |
Nitpick_HOL.register_codatatype @{typ "'a llist"} "" |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35312
diff
changeset
|
221 |
(map dest_Const [@{term LNil}, @{term LCons}]) |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35312
diff
changeset
|
222 |
*} |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35312
diff
changeset
|
223 |
|
33197 | 224 |
lemma "xs \<noteq> LCons a xs" |
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
225 |
nitpick [expect = genuine] |
33197 | 226 |
oops |
227 |
||
228 |
lemma "\<lbrakk>xs = LCons a xs; ys = iterates (\<lambda>b. a) b\<rbrakk> \<Longrightarrow> xs = ys" |
|
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
229 |
nitpick [verbose, expect = genuine] |
33197 | 230 |
oops |
231 |
||
232 |
lemma "\<lbrakk>xs = LCons a xs; ys = LCons a ys\<rbrakk> \<Longrightarrow> xs = ys" |
|
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
233 |
nitpick [bisim_depth = -1, show_datatypes, expect = quasi_genuine] |
42959 | 234 |
nitpick [card = 1\<emdash>5, expect = none] |
33197 | 235 |
sorry |
236 |
||
45053 | 237 |
subsection {* 2.10. Boxing *} |
33197 | 238 |
|
239 |
datatype tm = Var nat | Lam tm | App tm tm |
|
240 |
||
241 |
primrec lift where |
|
242 |
"lift (Var j) k = Var (if j < k then j else j + 1)" | |
|
243 |
"lift (Lam t) k = Lam (lift t (k + 1))" | |
|
244 |
"lift (App t u) k = App (lift t k) (lift u k)" |
|
245 |
||
246 |
primrec loose where |
|
247 |
"loose (Var j) k = (j \<ge> k)" | |
|
248 |
"loose (Lam t) k = loose t (Suc k)" | |
|
249 |
"loose (App t u) k = (loose t k \<or> loose u k)" |
|
250 |
||
251 |
primrec subst\<^isub>1 where |
|
252 |
"subst\<^isub>1 \<sigma> (Var j) = \<sigma> j" | |
|
253 |
"subst\<^isub>1 \<sigma> (Lam t) = |
|
254 |
Lam (subst\<^isub>1 (\<lambda>n. case n of 0 \<Rightarrow> Var 0 | Suc m \<Rightarrow> lift (\<sigma> m) 1) t)" | |
|
255 |
"subst\<^isub>1 \<sigma> (App t u) = App (subst\<^isub>1 \<sigma> t) (subst\<^isub>1 \<sigma> u)" |
|
256 |
||
257 |
lemma "\<not> loose t 0 \<Longrightarrow> subst\<^isub>1 \<sigma> t = t" |
|
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
258 |
nitpick [verbose, expect = genuine] |
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
259 |
nitpick [eval = "subst\<^isub>1 \<sigma> t", expect = genuine] |
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
260 |
(* nitpick [dont_box, expect = unknown] *) |
33197 | 261 |
oops |
262 |
||
263 |
primrec subst\<^isub>2 where |
|
264 |
"subst\<^isub>2 \<sigma> (Var j) = \<sigma> j" | |
|
265 |
"subst\<^isub>2 \<sigma> (Lam t) = |
|
266 |
Lam (subst\<^isub>2 (\<lambda>n. case n of 0 \<Rightarrow> Var 0 | Suc m \<Rightarrow> lift (\<sigma> m) 0) t)" | |
|
267 |
"subst\<^isub>2 \<sigma> (App t u) = App (subst\<^isub>2 \<sigma> t) (subst\<^isub>2 \<sigma> u)" |
|
268 |
||
269 |
lemma "\<not> loose t 0 \<Longrightarrow> subst\<^isub>2 \<sigma> t = t" |
|
42959 | 270 |
nitpick [card = 1\<emdash>5, expect = none] |
33197 | 271 |
sorry |
272 |
||
45053 | 273 |
subsection {* 2.11. Scope Monotonicity *} |
33197 | 274 |
|
275 |
lemma "length xs = length ys \<Longrightarrow> rev (zip xs ys) = zip xs (rev ys)" |
|
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
276 |
nitpick [verbose, expect = genuine] |
33197 | 277 |
oops |
278 |
||
279 |
lemma "\<exists>g. \<forall>x\<Colon>'b. g (f x) = x \<Longrightarrow> \<forall>y\<Colon>'a. \<exists>x. y = f x" |
|
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
280 |
nitpick [mono, expect = none] |
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
281 |
nitpick [expect = genuine] |
33197 | 282 |
oops |
283 |
||
45053 | 284 |
subsection {* 2.12. Inductive Properties *} |
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
285 |
|
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
286 |
inductive_set reach where |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
287 |
"(4\<Colon>nat) \<in> reach" | |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
288 |
"n \<in> reach \<Longrightarrow> n < 4 \<Longrightarrow> 3 * n + 1 \<in> reach" | |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
289 |
"n \<in> reach \<Longrightarrow> n + 2 \<in> reach" |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
290 |
|
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
291 |
lemma "n \<in> reach \<Longrightarrow> 2 dvd n" |
38184 | 292 |
(* nitpick *) |
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
293 |
apply (induct set: reach) |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
294 |
apply auto |
42959 | 295 |
nitpick [card = 1\<emdash>4, bits = 1\<emdash>4, expect = none] |
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
296 |
apply (thin_tac "n \<in> reach") |
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
297 |
nitpick [expect = genuine] |
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
298 |
oops |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
299 |
|
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
300 |
lemma "n \<in> reach \<Longrightarrow> 2 dvd n \<and> n \<noteq> 0" |
38184 | 301 |
(* nitpick *) |
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
302 |
apply (induct set: reach) |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
303 |
apply auto |
42959 | 304 |
nitpick [card = 1\<emdash>4, bits = 1\<emdash>4, expect = none] |
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
305 |
apply (thin_tac "n \<in> reach") |
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
306 |
nitpick [expect = genuine] |
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
307 |
oops |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
308 |
|
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
309 |
lemma "n \<in> reach \<Longrightarrow> 2 dvd n \<and> n \<ge> 4" |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
310 |
by (induct set: reach) arith+ |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
311 |
|
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
312 |
datatype 'a bin_tree = Leaf 'a | Branch "'a bin_tree" "'a bin_tree" |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
313 |
|
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
314 |
primrec labels where |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
315 |
"labels (Leaf a) = {a}" | |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
316 |
"labels (Branch t u) = labels t \<union> labels u" |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
317 |
|
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
318 |
primrec swap where |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
319 |
"swap (Leaf c) a b = |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
320 |
(if c = a then Leaf b else if c = b then Leaf a else Leaf c)" | |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
321 |
"swap (Branch t u) a b = Branch (swap t a b) (swap u a b)" |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
322 |
|
35180
c57dba973391
more work on Nitpick's support for nonstandard models + fix in model reconstruction
blanchet
parents:
35078
diff
changeset
|
323 |
lemma "{a, b} \<subseteq> labels t \<Longrightarrow> labels (swap t a b) = labels t" |
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
324 |
(* nitpick *) |
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
325 |
proof (induct t) |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
326 |
case Leaf thus ?case by simp |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
327 |
next |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
328 |
case (Branch t u) thus ?case |
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
329 |
(* nitpick *) |
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
330 |
nitpick [non_std, show_all, expect = genuine] |
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
331 |
oops |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
332 |
|
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
333 |
lemma "labels (swap t a b) = |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
334 |
(if a \<in> labels t then |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
335 |
if b \<in> labels t then labels t else (labels t - {a}) \<union> {b} |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
336 |
else |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
337 |
if b \<in> labels t then (labels t - {b}) \<union> {a} else labels t)" |
35309
997aa3a3e4bb
catch IO errors in Nitpick's "kodkodi" invocation + shorten execution time of "Manual_Nits" example
blanchet
parents:
35284
diff
changeset
|
338 |
(* nitpick *) |
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
339 |
proof (induct t) |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
340 |
case Leaf thus ?case by simp |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
341 |
next |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
342 |
case (Branch t u) thus ?case |
42959 | 343 |
nitpick [non_std, card = 1\<emdash>4, expect = none] |
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
344 |
by auto |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
345 |
qed |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
346 |
|
45053 | 347 |
section {* 3. Case Studies *} |
33197 | 348 |
|
36268 | 349 |
nitpick_params [max_potential = 0] |
33197 | 350 |
|
45053 | 351 |
subsection {* 3.1. A Context-Free Grammar *} |
33197 | 352 |
|
353 |
datatype alphabet = a | b |
|
354 |
||
355 |
inductive_set S\<^isub>1 and A\<^isub>1 and B\<^isub>1 where |
|
356 |
"[] \<in> S\<^isub>1" |
|
357 |
| "w \<in> A\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1" |
|
358 |
| "w \<in> B\<^isub>1 \<Longrightarrow> a # w \<in> S\<^isub>1" |
|
359 |
| "w \<in> S\<^isub>1 \<Longrightarrow> a # w \<in> A\<^isub>1" |
|
360 |
| "w \<in> S\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1" |
|
361 |
| "\<lbrakk>v \<in> B\<^isub>1; v \<in> B\<^isub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>1" |
|
362 |
||
363 |
theorem S\<^isub>1_sound: |
|
364 |
"w \<in> S\<^isub>1 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]" |
|
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
365 |
nitpick [expect = genuine] |
33197 | 366 |
oops |
367 |
||
368 |
inductive_set S\<^isub>2 and A\<^isub>2 and B\<^isub>2 where |
|
369 |
"[] \<in> S\<^isub>2" |
|
370 |
| "w \<in> A\<^isub>2 \<Longrightarrow> b # w \<in> S\<^isub>2" |
|
371 |
| "w \<in> B\<^isub>2 \<Longrightarrow> a # w \<in> S\<^isub>2" |
|
372 |
| "w \<in> S\<^isub>2 \<Longrightarrow> a # w \<in> A\<^isub>2" |
|
373 |
| "w \<in> S\<^isub>2 \<Longrightarrow> b # w \<in> B\<^isub>2" |
|
374 |
| "\<lbrakk>v \<in> B\<^isub>2; v \<in> B\<^isub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>2" |
|
375 |
||
376 |
theorem S\<^isub>2_sound: |
|
377 |
"w \<in> S\<^isub>2 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]" |
|
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
378 |
nitpick [expect = genuine] |
33197 | 379 |
oops |
380 |
||
381 |
inductive_set S\<^isub>3 and A\<^isub>3 and B\<^isub>3 where |
|
382 |
"[] \<in> S\<^isub>3" |
|
383 |
| "w \<in> A\<^isub>3 \<Longrightarrow> b # w \<in> S\<^isub>3" |
|
384 |
| "w \<in> B\<^isub>3 \<Longrightarrow> a # w \<in> S\<^isub>3" |
|
385 |
| "w \<in> S\<^isub>3 \<Longrightarrow> a # w \<in> A\<^isub>3" |
|
386 |
| "w \<in> S\<^isub>3 \<Longrightarrow> b # w \<in> B\<^isub>3" |
|
387 |
| "\<lbrakk>v \<in> B\<^isub>3; w \<in> B\<^isub>3\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>3" |
|
388 |
||
389 |
theorem S\<^isub>3_sound: |
|
390 |
"w \<in> S\<^isub>3 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]" |
|
42959 | 391 |
nitpick [card = 1\<emdash>5, expect = none] |
33197 | 392 |
sorry |
393 |
||
394 |
theorem S\<^isub>3_complete: |
|
395 |
"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^isub>3" |
|
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
396 |
nitpick [expect = genuine] |
33197 | 397 |
oops |
398 |
||
399 |
inductive_set S\<^isub>4 and A\<^isub>4 and B\<^isub>4 where |
|
400 |
"[] \<in> S\<^isub>4" |
|
401 |
| "w \<in> A\<^isub>4 \<Longrightarrow> b # w \<in> S\<^isub>4" |
|
402 |
| "w \<in> B\<^isub>4 \<Longrightarrow> a # w \<in> S\<^isub>4" |
|
403 |
| "w \<in> S\<^isub>4 \<Longrightarrow> a # w \<in> A\<^isub>4" |
|
404 |
| "\<lbrakk>v \<in> A\<^isub>4; w \<in> A\<^isub>4\<rbrakk> \<Longrightarrow> b # v @ w \<in> A\<^isub>4" |
|
405 |
| "w \<in> S\<^isub>4 \<Longrightarrow> b # w \<in> B\<^isub>4" |
|
406 |
| "\<lbrakk>v \<in> B\<^isub>4; w \<in> B\<^isub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>4" |
|
407 |
||
408 |
theorem S\<^isub>4_sound: |
|
409 |
"w \<in> S\<^isub>4 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]" |
|
42959 | 410 |
nitpick [card = 1\<emdash>5, expect = none] |
33197 | 411 |
sorry |
412 |
||
413 |
theorem S\<^isub>4_complete: |
|
414 |
"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^isub>4" |
|
42959 | 415 |
nitpick [card = 1\<emdash>5, expect = none] |
33197 | 416 |
sorry |
417 |
||
418 |
theorem S\<^isub>4_A\<^isub>4_B\<^isub>4_sound_and_complete: |
|
419 |
"w \<in> S\<^isub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]" |
|
420 |
"w \<in> A\<^isub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] + 1" |
|
421 |
"w \<in> B\<^isub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = b] = length [x \<leftarrow> w. x = a] + 1" |
|
42959 | 422 |
nitpick [card = 1\<emdash>5, expect = none] |
33197 | 423 |
sorry |
424 |
||
45053 | 425 |
subsection {* 3.2. AA Trees *} |
33197 | 426 |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34126
diff
changeset
|
427 |
datatype 'a aa_tree = \<Lambda> | N "'a\<Colon>linorder" nat "'a aa_tree" "'a aa_tree" |
33197 | 428 |
|
429 |
primrec data where |
|
430 |
"data \<Lambda> = undefined" | |
|
431 |
"data (N x _ _ _) = x" |
|
432 |
||
433 |
primrec dataset where |
|
434 |
"dataset \<Lambda> = {}" | |
|
435 |
"dataset (N x _ t u) = {x} \<union> dataset t \<union> dataset u" |
|
436 |
||
437 |
primrec level where |
|
438 |
"level \<Lambda> = 0" | |
|
439 |
"level (N _ k _ _) = k" |
|
440 |
||
441 |
primrec left where |
|
442 |
"left \<Lambda> = \<Lambda>" | |
|
443 |
"left (N _ _ t\<^isub>1 _) = t\<^isub>1" |
|
444 |
||
445 |
primrec right where |
|
446 |
"right \<Lambda> = \<Lambda>" | |
|
447 |
"right (N _ _ _ t\<^isub>2) = t\<^isub>2" |
|
448 |
||
449 |
fun wf where |
|
450 |
"wf \<Lambda> = True" | |
|
451 |
"wf (N _ k t u) = |
|
452 |
(if t = \<Lambda> then |
|
453 |
k = 1 \<and> (u = \<Lambda> \<or> (level u = 1 \<and> left u = \<Lambda> \<and> right u = \<Lambda>)) |
|
454 |
else |
|
455 |
wf t \<and> wf u \<and> u \<noteq> \<Lambda> \<and> level t < k \<and> level u \<le> k \<and> level (right u) < k)" |
|
456 |
||
457 |
fun skew where |
|
458 |
"skew \<Lambda> = \<Lambda>" | |
|
459 |
"skew (N x k t u) = |
|
460 |
(if t \<noteq> \<Lambda> \<and> k = level t then |
|
461 |
N (data t) k (left t) (N x k (right t) u) |
|
462 |
else |
|
463 |
N x k t u)" |
|
464 |
||
465 |
fun split where |
|
466 |
"split \<Lambda> = \<Lambda>" | |
|
467 |
"split (N x k t u) = |
|
468 |
(if u \<noteq> \<Lambda> \<and> k = level (right u) then |
|
469 |
N (data u) (Suc k) (N x k t (left u)) (right u) |
|
470 |
else |
|
471 |
N x k t u)" |
|
472 |
||
473 |
theorem dataset_skew_split: |
|
474 |
"dataset (skew t) = dataset t" |
|
475 |
"dataset (split t) = dataset t" |
|
42959 | 476 |
nitpick [card = 1\<emdash>5, expect = none] |
33197 | 477 |
sorry |
478 |
||
479 |
theorem wf_skew_split: |
|
480 |
"wf t \<Longrightarrow> skew t = t" |
|
481 |
"wf t \<Longrightarrow> split t = t" |
|
42959 | 482 |
nitpick [card = 1\<emdash>5, expect = none] |
33197 | 483 |
sorry |
484 |
||
485 |
primrec insort\<^isub>1 where |
|
486 |
"insort\<^isub>1 \<Lambda> x = N x 1 \<Lambda> \<Lambda>" | |
|
487 |
"insort\<^isub>1 (N y k t u) x = |
|
488 |
(* (split \<circ> skew) *) (N y k (if x < y then insort\<^isub>1 t x else t) |
|
489 |
(if x > y then insort\<^isub>1 u x else u))" |
|
490 |
||
491 |
theorem wf_insort\<^isub>1: "wf t \<Longrightarrow> wf (insort\<^isub>1 t x)" |
|
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
492 |
nitpick [expect = genuine] |
33197 | 493 |
oops |
494 |
||
495 |
theorem wf_insort\<^isub>1_nat: "wf t \<Longrightarrow> wf (insort\<^isub>1 t (x\<Colon>nat))" |
|
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
496 |
nitpick [eval = "insort\<^isub>1 t x", expect = genuine] |
33197 | 497 |
oops |
498 |
||
499 |
primrec insort\<^isub>2 where |
|
500 |
"insort\<^isub>2 \<Lambda> x = N x 1 \<Lambda> \<Lambda>" | |
|
501 |
"insort\<^isub>2 (N y k t u) x = |
|
502 |
(split \<circ> skew) (N y k (if x < y then insort\<^isub>2 t x else t) |
|
503 |
(if x > y then insort\<^isub>2 u x else u))" |
|
504 |
||
505 |
theorem wf_insort\<^isub>2: "wf t \<Longrightarrow> wf (insort\<^isub>2 t x)" |
|
42959 | 506 |
nitpick [card = 1\<emdash>5, expect = none] |
33197 | 507 |
sorry |
508 |
||
509 |
theorem dataset_insort\<^isub>2: "dataset (insort\<^isub>2 t x) = {x} \<union> dataset t" |
|
42959 | 510 |
nitpick [card = 1\<emdash>5, expect = none] |
33197 | 511 |
sorry |
512 |
||
513 |
end |