author | wenzelm |
Tue, 11 May 2021 20:19:07 +0200 | |
changeset 73675 | 6c56f2ebe157 |
parent 67613 | ce654b0e6d69 |
permissions | -rw-r--r-- |
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
1 |
theory Hotel_Example_Prolog |
41956 | 2 |
imports |
3 |
Hotel_Example |
|
66453
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
63167
diff
changeset
|
4 |
"HOL-Library.Predicate_Compile_Alternative_Defs" |
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
63167
diff
changeset
|
5 |
"HOL-Library.Code_Prolog" |
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
6 |
begin |
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
7 |
|
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
8 |
declare Let_def[code_pred_inline] |
45970
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
9 |
(* |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
10 |
thm hotel_def |
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
11 |
lemma [code_pred_inline]: "insert == (%y A x. y = x | A x)" |
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
12 |
by (auto simp add: insert_iff[unfolded mem_def] fun_eq_iff intro!: eq_reflection) |
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
13 |
|
67399 | 14 |
lemma [code_pred_inline]: "(-) == (%A B x. A x \<and> \<not> B x)" |
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
15 |
by (auto simp add: Diff_iff[unfolded mem_def] fun_eq_iff intro!: eq_reflection) |
45970
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
16 |
*) |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
17 |
|
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
18 |
definition issuedp :: "event list => key => bool" |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
19 |
where |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
20 |
"issuedp evs k = (k \<in> issued evs)" |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
21 |
|
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
22 |
lemma [code_pred_def]: |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
23 |
"issuedp [] Key0 = True" |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
24 |
"issuedp (e # s) k = (issuedp s k \<or> (case e of Check_in g r (k1, k2) => k = k2 | _ => False))" |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
25 |
unfolding issuedp_def issued.simps initk_def |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
26 |
by (auto split: event.split) |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
27 |
|
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
28 |
definition cardsp |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
29 |
where |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
30 |
"cardsp s g k = (k \<in> cards s g)" |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
31 |
|
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
32 |
lemma [code_pred_def]: |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
33 |
"cardsp [] g k = False" |
52666 | 34 |
"cardsp (e # s) g k = |
35 |
(let C = cardsp s g |
|
36 |
in case e of Check_in g' r c => if g' = g then k = c \<or> C k else C k | _ => C k)" |
|
46905 | 37 |
unfolding cardsp_def [abs_def] cards.simps by (auto simp add: Let_def split: event.split) |
45970
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
38 |
|
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
39 |
definition |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
40 |
"isinp evs r g = (g \<in> isin evs r)" |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
41 |
|
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
42 |
lemma [code_pred_def]: |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
43 |
"isinp [] r g = False" |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
44 |
"isinp (e # s) r g = |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
45 |
(let G = isinp s r |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
46 |
in case e of Check_in g' r c => G g |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
47 |
| Enter g' r' c => if r' = r then g = g' \<or> G g else G g |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
48 |
| Exit g' r' => if r' = r then ((g \<noteq> g') \<and> G g) else G g)" |
46905 | 49 |
unfolding isinp_def [abs_def] isin.simps by (auto simp add: Let_def split: event.split) |
45970
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
50 |
|
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
51 |
declare hotel.simps(1)[code_pred_def] |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
52 |
lemma [code_pred_def]: |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
53 |
"hotel (e # s) = |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
54 |
(hotel s & (case e of Check_in g r (k, k') => k = currk s r & \<not> issuedp s k' |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
55 |
| Enter g r (k, k') => cardsp s g (k, k') & (roomk s r = k \<or> roomk s r = k') |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
56 |
| Exit g r => isinp s r g))" |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
57 |
unfolding hotel.simps issuedp_def cardsp_def isinp_def |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
58 |
by (auto split: event.split) |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
59 |
|
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
60 |
declare List.member_rec[code_pred_def] |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
61 |
|
67613 | 62 |
lemma [code_pred_def]: "no_Check_in s r = (\<not> (\<exists>g c. List.member s (Check_in g r c)))" |
45970
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
63 |
unfolding no_Check_in_def member_def by auto |
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
64 |
|
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
65 |
lemma [code_pred_def]: "feels_safe s r = |
67613 | 66 |
(\<exists>s\<^sub>1 s\<^sub>2 s\<^sub>3 g c c'. |
45970
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
67 |
s = |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52666
diff
changeset
|
68 |
s\<^sub>3 @ |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52666
diff
changeset
|
69 |
[Enter g r c] @ s\<^sub>2 @ [Check_in g r c'] @ s\<^sub>1 & |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52666
diff
changeset
|
70 |
no_Check_in (s\<^sub>3 @ s\<^sub>2) r & |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52666
diff
changeset
|
71 |
(\<not> (\<exists> g'. isinp (s\<^sub>2 @ [Check_in g r c] @ s\<^sub>1) r g')))" |
45970
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
72 |
unfolding feels_safe_def isinp_def by auto |
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
73 |
|
63167 | 74 |
setup \<open>Code_Prolog.map_code_options (K |
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
75 |
{ensure_groundness = true, |
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
76 |
limit_globally = NONE, |
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
77 |
limited_types = [], |
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
78 |
limited_predicates = [], |
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
79 |
replacing = [], |
63167 | 80 |
manual_reorder = []})\<close> |
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
81 |
|
55447 | 82 |
values_prolog 40 "{s. hotel s}" |
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
83 |
|
63167 | 84 |
setup \<open> |
52666 | 85 |
Context.theory_map |
86 |
(Quickcheck.add_tester ("prolog", (Code_Prolog.active, Code_Prolog.test_goals))) |
|
63167 | 87 |
\<close> |
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
88 |
|
45970
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
89 |
lemma "\<lbrakk> hotel s; isinp s r g \<rbrakk> \<Longrightarrow> owns s r = Some g" |
40924 | 90 |
quickcheck[tester = prolog, iterations = 1, expect = counterexample] |
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
91 |
oops |
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
92 |
|
63167 | 93 |
section \<open>Manual setup to find the counterexample\<close> |
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
94 |
|
63167 | 95 |
setup \<open>Code_Prolog.map_code_options (K |
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
96 |
{ensure_groundness = true, |
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
97 |
limit_globally = NONE, |
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
98 |
limited_types = [], |
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
99 |
limited_predicates = [(["hotel"], 4)], |
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
100 |
replacing = [(("hotel", "limited_hotel"), "quickcheck")], |
63167 | 101 |
manual_reorder = []})\<close> |
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
102 |
|
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
103 |
lemma |
45970
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
104 |
"hotel s ==> feels_safe s r ==> isinp s r g ==> owns s r = Some g" |
40924 | 105 |
quickcheck[tester = prolog, iterations = 1, expect = no_counterexample] |
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
106 |
oops |
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
107 |
|
63167 | 108 |
setup \<open>Code_Prolog.map_code_options (K |
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
109 |
{ensure_groundness = true, |
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
110 |
limit_globally = NONE, |
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
111 |
limited_types = [], |
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
112 |
limited_predicates = [(["hotel"], 5)], |
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
113 |
replacing = [(("hotel", "limited_hotel"), "quickcheck")], |
63167 | 114 |
manual_reorder = []})\<close> |
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
115 |
|
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
116 |
lemma |
45970
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
117 |
"hotel s ==> feels_safe s r ==> isinp s r g ==> owns s r = Some g" |
40924 | 118 |
quickcheck[tester = prolog, iterations = 1, expect = counterexample] |
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
119 |
oops |
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
120 |
|
63167 | 121 |
section \<open>Using a global limit for limiting the execution\<close> |
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
122 |
|
63167 | 123 |
text \<open>A global depth limit of 10 does not suffice to find the counterexample.\<close> |
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
124 |
|
63167 | 125 |
setup \<open> |
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
126 |
Code_Prolog.map_code_options (K |
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
127 |
{ensure_groundness = true, |
45970
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
128 |
limit_globally = SOME 10, |
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
129 |
limited_types = [], |
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
130 |
limited_predicates = [], |
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
131 |
replacing = [], |
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
132 |
manual_reorder = []}) |
63167 | 133 |
\<close> |
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
134 |
|
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
135 |
lemma |
45970
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
136 |
"hotel s ==> feels_safe s r ==> isinp s r g ==> owns s r = Some g" |
40924 | 137 |
quickcheck[tester = prolog, iterations = 1, expect = no_counterexample] |
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
138 |
oops |
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
139 |
|
63167 | 140 |
text \<open>But a global depth limit of 11 does.\<close> |
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
141 |
|
63167 | 142 |
setup \<open> |
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
143 |
Code_Prolog.map_code_options (K |
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
144 |
{ensure_groundness = true, |
45970
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
145 |
limit_globally = SOME 11, |
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
146 |
limited_types = [], |
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
147 |
limited_predicates = [], |
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
148 |
replacing = [], |
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
149 |
manual_reorder = []}) |
63167 | 150 |
\<close> |
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
151 |
|
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
152 |
lemma |
45970
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
43885
diff
changeset
|
153 |
"hotel s ==> feels_safe s r ==> isinp s r g ==> owns s r = Some g" |
40924 | 154 |
quickcheck[tester = prolog, iterations = 1, expect = counterexample] |
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
155 |
oops |
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff
changeset
|
156 |
|
62390 | 157 |
end |