author | wenzelm |
Sun, 21 Aug 2022 15:00:14 +0200 | |
changeset 75952 | 864b10457a7d |
parent 74641 | 6f801e1073fa |
permissions | -rw-r--r-- |
33197 | 1 |
(* Title: HOL/Nitpick_Examples/Manual_Nits.thy |
2 |
Author: Jasmin Blanchette, TU Muenchen |
|
55888 | 3 |
Copyright 2009-2014 |
33197 | 4 |
|
5 |
Examples from the Nitpick manual. |
|
6 |
*) |
|
7 |
||
63167 | 8 |
section \<open>Examples from the Nitpick Manual\<close> |
33197 | 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 |
66453
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
63167
diff
changeset
|
15 |
imports HOL.Real "HOL-Library.Quotient_Product" |
33197 | 16 |
begin |
17 |
||
63167 | 18 |
section \<open>2. First Steps\<close> |
33197 | 19 |
|
74641
6f801e1073fa
prefer "sat_solver = MiniSat", to make examples work uniformly on all platforms;
wenzelm
parents:
74399
diff
changeset
|
20 |
nitpick_params [sat_solver = MiniSat, max_threads = 1, timeout = 240] |
33197 | 21 |
|
55888 | 22 |
|
63167 | 23 |
subsection \<open>2.1. Propositional Logic\<close> |
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 |
||
55888 | 32 |
|
63167 | 33 |
subsection \<open>2.2. Type Variables\<close> |
33197 | 34 |
|
46104 | 35 |
lemma "x \<in> A \<Longrightarrow> (THE y. y \<in> A) \<in> A" |
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
|
36 |
nitpick [verbose, expect = genuine] |
33197 | 37 |
oops |
38 |
||
55888 | 39 |
|
63167 | 40 |
subsection \<open>2.3. Constants\<close> |
33197 | 41 |
|
46104 | 42 |
lemma "x \<in> A \<Longrightarrow> (THE y. y \<in> A) \<in> A" |
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
|
43 |
nitpick [show_consts, expect = genuine] |
39362
ee65900bfced
adapt examples to latest Nitpick changes + speed them up a little bit
blanchet
parents:
39302
diff
changeset
|
44 |
nitpick [dont_specialize, show_consts, expect = genuine] |
33197 | 45 |
oops |
46 |
||
46104 | 47 |
lemma "\<exists>!x. x \<in> A \<Longrightarrow> (THE y. y \<in> A) \<in> A" |
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
|
48 |
nitpick [expect = none] |
55893
aed17a173d16
fixed handling of 'case_prod' and other 'case' functions for interpreted types
blanchet
parents:
55889
diff
changeset
|
49 |
nitpick [card 'a = 1-50, expect = none] |
33197 | 50 |
(* sledgehammer *) |
46104 | 51 |
by (metis the_equality) |
33197 | 52 |
|
55888 | 53 |
|
63167 | 54 |
subsection \<open>2.4. Skolemization\<close> |
33197 | 55 |
|
56 |
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
|
57 |
nitpick [expect = genuine] |
33197 | 58 |
oops |
59 |
||
60 |
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
|
61 |
nitpick [expect = genuine] |
33197 | 62 |
oops |
63 |
||
64 |
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
|
65 |
nitpick [expect = genuine] |
33197 | 66 |
oops |
67 |
||
55888 | 68 |
|
63167 | 69 |
subsection \<open>2.5. Natural Numbers and Integers\<close> |
33197 | 70 |
|
61076 | 71 |
lemma "\<lbrakk>i \<le> j; n \<le> (m::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
|
72 |
nitpick [expect = genuine] |
46104 | 73 |
nitpick [binary_ints, bits = 16, expect = genuine] |
33197 | 74 |
oops |
75 |
||
76 |
lemma "\<forall>n. Suc n \<noteq> n \<Longrightarrow> P" |
|
60310 | 77 |
nitpick [card nat = 100, expect = potential] |
33197 | 78 |
oops |
79 |
||
80 |
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
|
81 |
nitpick [expect = none] |
33197 | 82 |
oops |
83 |
||
67399 | 84 |
lemma "P ((+)::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
|
85 |
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
|
86 |
nitpick [card nat = 2, expect = none] |
33197 | 87 |
oops |
88 |
||
55888 | 89 |
|
63167 | 90 |
subsection \<open>2.6. Inductive Datatypes\<close> |
33197 | 91 |
|
92 |
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
|
93 |
nitpick [expect = genuine] |
55889 | 94 |
nitpick [show_consts, show_types, expect = genuine] |
33197 | 95 |
oops |
96 |
||
97 |
lemma "\<lbrakk>length xs = 1; length ys = 1\<rbrakk> \<Longrightarrow> xs = ys" |
|
55889 | 98 |
nitpick [show_types, expect = genuine] |
33197 | 99 |
oops |
100 |
||
55888 | 101 |
|
63167 | 102 |
subsection \<open>2.7. Typedefs, Records, Rationals, and Reals\<close> |
33197 | 103 |
|
61076 | 104 |
definition "three = {0::nat, 1, 2}" |
49834 | 105 |
typedef three = three |
49812
e3945ddcb9aa
eliminated some remaining uses of typedef with implicit set definition;
wenzelm
parents:
48812
diff
changeset
|
106 |
unfolding three_def by blast |
33197 | 107 |
|
108 |
definition A :: three where "A \<equiv> Abs_three 0" |
|
109 |
definition B :: three where "B \<equiv> Abs_three 1" |
|
110 |
definition C :: three where "C \<equiv> Abs_three 2" |
|
111 |
||
46104 | 112 |
lemma "\<lbrakk>A \<in> X; B \<in> X\<rbrakk> \<Longrightarrow> c \<in> X" |
55889 | 113 |
nitpick [show_types, expect = genuine] |
33197 | 114 |
oops |
115 |
||
35284
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents:
35185
diff
changeset
|
116 |
fun my_int_rel where |
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents:
35185
diff
changeset
|
117 |
"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
|
118 |
|
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents:
35185
diff
changeset
|
119 |
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
|
120 |
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
|
121 |
|
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents:
35185
diff
changeset
|
122 |
definition add_raw where |
61076 | 123 |
"add_raw \<equiv> \<lambda>(x, y) (u, v). (x + (u::nat), y + (v::nat))" |
35284
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents:
35185
diff
changeset
|
124 |
|
61076 | 125 |
quotient_definition "add::my_int \<Rightarrow> my_int \<Rightarrow> my_int" is add_raw |
47092 | 126 |
unfolding add_raw_def by auto |
35284
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents:
35185
diff
changeset
|
127 |
|
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents:
35185
diff
changeset
|
128 |
lemma "add x y = add x x" |
55889 | 129 |
nitpick [show_types, expect = genuine] |
35284
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents:
35185
diff
changeset
|
130 |
oops |
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents:
35185
diff
changeset
|
131 |
|
63167 | 132 |
ML \<open> |
35712
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
133 |
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
|
134 |
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
|
135 |
- 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
|
136 |
| my_int_postproc _ _ _ _ t = t |
63167 | 137 |
\<close> |
35711
548d3f16404b
added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents:
35710
diff
changeset
|
138 |
|
63167 | 139 |
declaration \<open> |
74399 | 140 |
Nitpick_Model.register_term_postprocessor \<^Type>\<open>my_int\<close> my_int_postproc |
63167 | 141 |
\<close> |
35711
548d3f16404b
added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents:
35710
diff
changeset
|
142 |
|
548d3f16404b
added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents:
35710
diff
changeset
|
143 |
lemma "add x y = add x x" |
55889 | 144 |
nitpick [show_types] |
35711
548d3f16404b
added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents:
35710
diff
changeset
|
145 |
oops |
548d3f16404b
added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents:
35710
diff
changeset
|
146 |
|
33197 | 147 |
record point = |
148 |
Xcoord :: int |
|
149 |
Ycoord :: int |
|
150 |
||
61076 | 151 |
lemma "Xcoord (p::point) = Xcoord (q::point)" |
55889 | 152 |
nitpick [show_types, expect = genuine] |
33197 | 153 |
oops |
154 |
||
61076 | 155 |
lemma "4 * x + 3 * (y::real) \<noteq> 1 / 2" |
55889 | 156 |
nitpick [show_types, expect = genuine] |
33197 | 157 |
oops |
158 |
||
55888 | 159 |
|
63167 | 160 |
subsection \<open>2.8. Inductive and Coinductive Predicates\<close> |
33197 | 161 |
|
162 |
inductive even where |
|
163 |
"even 0" | |
|
164 |
"even n \<Longrightarrow> even (Suc (Suc n))" |
|
165 |
||
166 |
lemma "\<exists>n. even n \<and> even (Suc n)" |
|
35710 | 167 |
nitpick [card nat = 50, unary_ints, verbose, expect = potential] |
33197 | 168 |
oops |
169 |
||
35710 | 170 |
lemma "\<exists>n \<le> 49. even n \<and> even (Suc n)" |
38184 | 171 |
nitpick [card nat = 50, unary_ints, expect = genuine] |
33197 | 172 |
oops |
173 |
||
174 |
inductive even' where |
|
61076 | 175 |
"even' (0::nat)" | |
33197 | 176 |
"even' 2" | |
177 |
"\<lbrakk>even' m; even' n\<rbrakk> \<Longrightarrow> even' (m + n)" |
|
178 |
||
179 |
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
|
180 |
nitpick [card nat = 10, unary_ints, verbose, show_consts, expect = genuine] |
33197 | 181 |
oops |
182 |
||
183 |
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
|
184 |
nitpick [card nat = 10, show_consts, expect = genuine] |
33197 | 185 |
oops |
186 |
||
187 |
coinductive nats where |
|
61076 | 188 |
"nats (x::nat) \<Longrightarrow> nats x" |
33197 | 189 |
|
45970
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
45694
diff
changeset
|
190 |
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
|
191 |
nitpick [card nat = 10, show_consts, expect = genuine] |
33197 | 192 |
oops |
193 |
||
194 |
inductive odd where |
|
195 |
"odd 1" | |
|
196 |
"\<lbrakk>odd m; even n\<rbrakk> \<Longrightarrow> odd (m + n)" |
|
197 |
||
198 |
lemma "odd n \<Longrightarrow> odd (n - 2)" |
|
46105 | 199 |
nitpick [card nat = 4, show_consts, expect = genuine] |
33197 | 200 |
oops |
201 |
||
55888 | 202 |
|
63167 | 203 |
subsection \<open>2.9. Coinductive Datatypes\<close> |
33197 | 204 |
|
53808 | 205 |
codatatype 'a llist = LNil | LCons 'a "'a llist" |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35312
diff
changeset
|
206 |
|
53827 | 207 |
primcorec iterates where |
208 |
"iterates f a = LCons a (iterates f (f a))" |
|
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35312
diff
changeset
|
209 |
|
33197 | 210 |
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
|
211 |
nitpick [expect = genuine] |
33197 | 212 |
oops |
213 |
||
214 |
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
|
215 |
nitpick [verbose, expect = genuine] |
33197 | 216 |
oops |
217 |
||
218 |
lemma "\<lbrakk>xs = LCons a xs; ys = LCons a ys\<rbrakk> \<Longrightarrow> xs = ys" |
|
55889 | 219 |
nitpick [bisim_depth = -1, show_types, expect = quasi_genuine] |
55893
aed17a173d16
fixed handling of 'case_prod' and other 'case' functions for interpreted types
blanchet
parents:
55889
diff
changeset
|
220 |
nitpick [card = 1-5, expect = none] |
33197 | 221 |
sorry |
222 |
||
55888 | 223 |
|
63167 | 224 |
subsection \<open>2.10. Boxing\<close> |
33197 | 225 |
|
58310 | 226 |
datatype tm = Var nat | Lam tm | App tm tm |
33197 | 227 |
|
228 |
primrec lift where |
|
229 |
"lift (Var j) k = Var (if j < k then j else j + 1)" | |
|
230 |
"lift (Lam t) k = Lam (lift t (k + 1))" | |
|
231 |
"lift (App t u) k = App (lift t k) (lift u k)" |
|
232 |
||
233 |
primrec loose where |
|
234 |
"loose (Var j) k = (j \<ge> k)" | |
|
235 |
"loose (Lam t) k = loose t (Suc k)" | |
|
236 |
"loose (App t u) k = (loose t k \<or> loose u k)" |
|
237 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
238 |
primrec subst\<^sub>1 where |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
239 |
"subst\<^sub>1 \<sigma> (Var j) = \<sigma> j" | |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
240 |
"subst\<^sub>1 \<sigma> (Lam t) = |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
241 |
Lam (subst\<^sub>1 (\<lambda>n. case n of 0 \<Rightarrow> Var 0 | Suc m \<Rightarrow> lift (\<sigma> m) 1) t)" | |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
242 |
"subst\<^sub>1 \<sigma> (App t u) = App (subst\<^sub>1 \<sigma> t) (subst\<^sub>1 \<sigma> u)" |
33197 | 243 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
244 |
lemma "\<not> loose t 0 \<Longrightarrow> subst\<^sub>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
|
245 |
nitpick [verbose, expect = genuine] |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
246 |
nitpick [eval = "subst\<^sub>1 \<sigma> t", expect = genuine] |
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
|
247 |
(* nitpick [dont_box, expect = unknown] *) |
33197 | 248 |
oops |
249 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
250 |
primrec subst\<^sub>2 where |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
251 |
"subst\<^sub>2 \<sigma> (Var j) = \<sigma> j" | |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
252 |
"subst\<^sub>2 \<sigma> (Lam t) = |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
253 |
Lam (subst\<^sub>2 (\<lambda>n. case n of 0 \<Rightarrow> Var 0 | Suc m \<Rightarrow> lift (\<sigma> m) 0) t)" | |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
254 |
"subst\<^sub>2 \<sigma> (App t u) = App (subst\<^sub>2 \<sigma> t) (subst\<^sub>2 \<sigma> u)" |
33197 | 255 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
256 |
lemma "\<not> loose t 0 \<Longrightarrow> subst\<^sub>2 \<sigma> t = t" |
55893
aed17a173d16
fixed handling of 'case_prod' and other 'case' functions for interpreted types
blanchet
parents:
55889
diff
changeset
|
257 |
nitpick [card = 1-5, expect = none] |
33197 | 258 |
sorry |
259 |
||
55888 | 260 |
|
63167 | 261 |
subsection \<open>2.11. Scope Monotonicity\<close> |
33197 | 262 |
|
263 |
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
|
264 |
nitpick [verbose, expect = genuine] |
33197 | 265 |
oops |
266 |
||
61076 | 267 |
lemma "\<exists>g. \<forall>x::'b. g (f x) = x \<Longrightarrow> \<forall>y::'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
|
268 |
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
|
269 |
nitpick [expect = genuine] |
33197 | 270 |
oops |
271 |
||
55888 | 272 |
|
63167 | 273 |
subsection \<open>2.12. Inductive Properties\<close> |
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
|
274 |
|
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
|
275 |
inductive_set reach where |
61076 | 276 |
"(4::nat) \<in> reach" | |
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
|
277 |
"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
|
278 |
"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
|
279 |
|
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
|
280 |
lemma "n \<in> reach \<Longrightarrow> 2 dvd n" |
38184 | 281 |
(* 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
|
282 |
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
|
283 |
apply auto |
55893
aed17a173d16
fixed handling of 'case_prod' and other 'case' functions for interpreted types
blanchet
parents:
55889
diff
changeset
|
284 |
nitpick [card = 1-4, bits = 1-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
|
285 |
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
|
286 |
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
|
287 |
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
|
288 |
|
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 |
lemma "n \<in> reach \<Longrightarrow> 2 dvd n \<and> n \<noteq> 0" |
38184 | 290 |
(* 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
|
291 |
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
|
292 |
apply auto |
55893
aed17a173d16
fixed handling of 'case_prod' and other 'case' functions for interpreted types
blanchet
parents:
55889
diff
changeset
|
293 |
nitpick [card = 1-4, bits = 1-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
|
294 |
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
|
295 |
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
|
296 |
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
|
297 |
|
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 |
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
|
299 |
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
|
300 |
|
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
|
301 |
|
63167 | 302 |
section \<open>3. Case Studies\<close> |
33197 | 303 |
|
36268 | 304 |
nitpick_params [max_potential = 0] |
33197 | 305 |
|
55888 | 306 |
|
63167 | 307 |
subsection \<open>3.1. A Context-Free Grammar\<close> |
33197 | 308 |
|
58310 | 309 |
datatype alphabet = a | b |
33197 | 310 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
311 |
inductive_set S\<^sub>1 and A\<^sub>1 and B\<^sub>1 where |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
312 |
"[] \<in> S\<^sub>1" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
313 |
| "w \<in> A\<^sub>1 \<Longrightarrow> b # w \<in> S\<^sub>1" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
314 |
| "w \<in> B\<^sub>1 \<Longrightarrow> a # w \<in> S\<^sub>1" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
315 |
| "w \<in> S\<^sub>1 \<Longrightarrow> a # w \<in> A\<^sub>1" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
316 |
| "w \<in> S\<^sub>1 \<Longrightarrow> b # w \<in> S\<^sub>1" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
317 |
| "\<lbrakk>v \<in> B\<^sub>1; v \<in> B\<^sub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>1" |
33197 | 318 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
319 |
theorem S\<^sub>1_sound: |
68386 | 320 |
"w \<in> S\<^sub>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
|
321 |
nitpick [expect = genuine] |
33197 | 322 |
oops |
323 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
324 |
inductive_set S\<^sub>2 and A\<^sub>2 and B\<^sub>2 where |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
325 |
"[] \<in> S\<^sub>2" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
326 |
| "w \<in> A\<^sub>2 \<Longrightarrow> b # w \<in> S\<^sub>2" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
327 |
| "w \<in> B\<^sub>2 \<Longrightarrow> a # w \<in> S\<^sub>2" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
328 |
| "w \<in> S\<^sub>2 \<Longrightarrow> a # w \<in> A\<^sub>2" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
329 |
| "w \<in> S\<^sub>2 \<Longrightarrow> b # w \<in> B\<^sub>2" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
330 |
| "\<lbrakk>v \<in> B\<^sub>2; v \<in> B\<^sub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>2" |
33197 | 331 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
332 |
theorem S\<^sub>2_sound: |
68386 | 333 |
"w \<in> S\<^sub>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
|
334 |
nitpick [expect = genuine] |
33197 | 335 |
oops |
336 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
337 |
inductive_set S\<^sub>3 and A\<^sub>3 and B\<^sub>3 where |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
338 |
"[] \<in> S\<^sub>3" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
339 |
| "w \<in> A\<^sub>3 \<Longrightarrow> b # w \<in> S\<^sub>3" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
340 |
| "w \<in> B\<^sub>3 \<Longrightarrow> a # w \<in> S\<^sub>3" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
341 |
| "w \<in> S\<^sub>3 \<Longrightarrow> a # w \<in> A\<^sub>3" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
342 |
| "w \<in> S\<^sub>3 \<Longrightarrow> b # w \<in> B\<^sub>3" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
343 |
| "\<lbrakk>v \<in> B\<^sub>3; w \<in> B\<^sub>3\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>3" |
33197 | 344 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
345 |
theorem S\<^sub>3_sound: |
68386 | 346 |
"w \<in> S\<^sub>3 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]" |
55893
aed17a173d16
fixed handling of 'case_prod' and other 'case' functions for interpreted types
blanchet
parents:
55889
diff
changeset
|
347 |
nitpick [card = 1-5, expect = none] |
33197 | 348 |
sorry |
349 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
350 |
theorem S\<^sub>3_complete: |
68386 | 351 |
"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^sub>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
|
352 |
nitpick [expect = genuine] |
33197 | 353 |
oops |
354 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
355 |
inductive_set S\<^sub>4 and A\<^sub>4 and B\<^sub>4 where |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
356 |
"[] \<in> S\<^sub>4" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
357 |
| "w \<in> A\<^sub>4 \<Longrightarrow> b # w \<in> S\<^sub>4" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
358 |
| "w \<in> B\<^sub>4 \<Longrightarrow> a # w \<in> S\<^sub>4" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
359 |
| "w \<in> S\<^sub>4 \<Longrightarrow> a # w \<in> A\<^sub>4" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
360 |
| "\<lbrakk>v \<in> A\<^sub>4; w \<in> A\<^sub>4\<rbrakk> \<Longrightarrow> b # v @ w \<in> A\<^sub>4" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
361 |
| "w \<in> S\<^sub>4 \<Longrightarrow> b # w \<in> B\<^sub>4" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
362 |
| "\<lbrakk>v \<in> B\<^sub>4; w \<in> B\<^sub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>4" |
33197 | 363 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
364 |
theorem S\<^sub>4_sound: |
68386 | 365 |
"w \<in> S\<^sub>4 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]" |
55893
aed17a173d16
fixed handling of 'case_prod' and other 'case' functions for interpreted types
blanchet
parents:
55889
diff
changeset
|
366 |
nitpick [card = 1-5, expect = none] |
33197 | 367 |
sorry |
368 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
369 |
theorem S\<^sub>4_complete: |
68386 | 370 |
"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^sub>4" |
55893
aed17a173d16
fixed handling of 'case_prod' and other 'case' functions for interpreted types
blanchet
parents:
55889
diff
changeset
|
371 |
nitpick [card = 1-5, expect = none] |
33197 | 372 |
sorry |
373 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
374 |
theorem S\<^sub>4_A\<^sub>4_B\<^sub>4_sound_and_complete: |
68386 | 375 |
"w \<in> S\<^sub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]" |
376 |
"w \<in> A\<^sub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] + 1" |
|
377 |
"w \<in> B\<^sub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = b] = length [x \<leftarrow> w. x = a] + 1" |
|
55893
aed17a173d16
fixed handling of 'case_prod' and other 'case' functions for interpreted types
blanchet
parents:
55889
diff
changeset
|
378 |
nitpick [card = 1-5, expect = none] |
33197 | 379 |
sorry |
380 |
||
55888 | 381 |
|
63167 | 382 |
subsection \<open>3.2. AA Trees\<close> |
33197 | 383 |
|
61076 | 384 |
datatype 'a aa_tree = \<Lambda> | N "'a::linorder" nat "'a aa_tree" "'a aa_tree" |
33197 | 385 |
|
386 |
primrec data where |
|
387 |
"data \<Lambda> = undefined" | |
|
388 |
"data (N x _ _ _) = x" |
|
389 |
||
390 |
primrec dataset where |
|
391 |
"dataset \<Lambda> = {}" | |
|
392 |
"dataset (N x _ t u) = {x} \<union> dataset t \<union> dataset u" |
|
393 |
||
394 |
primrec level where |
|
395 |
"level \<Lambda> = 0" | |
|
396 |
"level (N _ k _ _) = k" |
|
397 |
||
398 |
primrec left where |
|
399 |
"left \<Lambda> = \<Lambda>" | |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
400 |
"left (N _ _ t\<^sub>1 _) = t\<^sub>1" |
33197 | 401 |
|
402 |
primrec right where |
|
403 |
"right \<Lambda> = \<Lambda>" | |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
404 |
"right (N _ _ _ t\<^sub>2) = t\<^sub>2" |
33197 | 405 |
|
406 |
fun wf where |
|
407 |
"wf \<Lambda> = True" | |
|
408 |
"wf (N _ k t u) = |
|
409 |
(if t = \<Lambda> then |
|
410 |
k = 1 \<and> (u = \<Lambda> \<or> (level u = 1 \<and> left u = \<Lambda> \<and> right u = \<Lambda>)) |
|
411 |
else |
|
412 |
wf t \<and> wf u \<and> u \<noteq> \<Lambda> \<and> level t < k \<and> level u \<le> k \<and> level (right u) < k)" |
|
413 |
||
414 |
fun skew where |
|
415 |
"skew \<Lambda> = \<Lambda>" | |
|
416 |
"skew (N x k t u) = |
|
417 |
(if t \<noteq> \<Lambda> \<and> k = level t then |
|
418 |
N (data t) k (left t) (N x k (right t) u) |
|
419 |
else |
|
420 |
N x k t u)" |
|
421 |
||
422 |
fun split where |
|
423 |
"split \<Lambda> = \<Lambda>" | |
|
424 |
"split (N x k t u) = |
|
425 |
(if u \<noteq> \<Lambda> \<and> k = level (right u) then |
|
426 |
N (data u) (Suc k) (N x k t (left u)) (right u) |
|
427 |
else |
|
428 |
N x k t u)" |
|
429 |
||
430 |
theorem dataset_skew_split: |
|
431 |
"dataset (skew t) = dataset t" |
|
432 |
"dataset (split t) = dataset t" |
|
55893
aed17a173d16
fixed handling of 'case_prod' and other 'case' functions for interpreted types
blanchet
parents:
55889
diff
changeset
|
433 |
nitpick [card = 1-5, expect = none] |
33197 | 434 |
sorry |
435 |
||
436 |
theorem wf_skew_split: |
|
437 |
"wf t \<Longrightarrow> skew t = t" |
|
438 |
"wf t \<Longrightarrow> split t = t" |
|
55893
aed17a173d16
fixed handling of 'case_prod' and other 'case' functions for interpreted types
blanchet
parents:
55889
diff
changeset
|
439 |
nitpick [card = 1-5, expect = none] |
33197 | 440 |
sorry |
441 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
442 |
primrec insort\<^sub>1 where |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
443 |
"insort\<^sub>1 \<Lambda> x = N x 1 \<Lambda> \<Lambda>" | |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
444 |
"insort\<^sub>1 (N y k t u) x = |
67414 | 445 |
\<^cancel>\<open>(split \<circ> skew)\<close> (N y k (if x < y then insort\<^sub>1 t x else t) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
446 |
(if x > y then insort\<^sub>1 u x else u))" |
33197 | 447 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
448 |
theorem wf_insort\<^sub>1: "wf t \<Longrightarrow> wf (insort\<^sub>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
|
449 |
nitpick [expect = genuine] |
33197 | 450 |
oops |
451 |
||
61076 | 452 |
theorem wf_insort\<^sub>1_nat: "wf t \<Longrightarrow> wf (insort\<^sub>1 t (x::nat))" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
453 |
nitpick [eval = "insort\<^sub>1 t x", expect = genuine] |
33197 | 454 |
oops |
455 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
456 |
primrec insort\<^sub>2 where |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
457 |
"insort\<^sub>2 \<Lambda> x = N x 1 \<Lambda> \<Lambda>" | |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
458 |
"insort\<^sub>2 (N y k t u) x = |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
459 |
(split \<circ> skew) (N y k (if x < y then insort\<^sub>2 t x else t) |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
460 |
(if x > y then insort\<^sub>2 u x else u))" |
33197 | 461 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
462 |
theorem wf_insort\<^sub>2: "wf t \<Longrightarrow> wf (insort\<^sub>2 t x)" |
55893
aed17a173d16
fixed handling of 'case_prod' and other 'case' functions for interpreted types
blanchet
parents:
55889
diff
changeset
|
463 |
nitpick [card = 1-5, expect = none] |
33197 | 464 |
sorry |
465 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51523
diff
changeset
|
466 |
theorem dataset_insort\<^sub>2: "dataset (insort\<^sub>2 t x) = {x} \<union> dataset t" |
55893
aed17a173d16
fixed handling of 'case_prod' and other 'case' functions for interpreted types
blanchet
parents:
55889
diff
changeset
|
467 |
nitpick [card = 1-5, expect = none] |
33197 | 468 |
sorry |
469 |
||
470 |
end |