src/HOL/Predicate_Compile_Examples/Hotel_Example_Prolog.thy
author Gerwin Klein <gerwin.klein@nicta.com.au>
Mon, 09 Jul 2012 09:15:28 +1000
changeset 48210 08cb859c53cd
parent 46905 6b1c0a80a57a
child 52666 391913d17d15
permissions -rw-r--r--
removed switched-off atbroy102 from isatest; no cygwin test currently running, needs a working server somewhere
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
c15ef1b85035 modernized imports (untested!?);
wenzelm
parents: 41413
diff changeset
     2
imports
c15ef1b85035 modernized imports (untested!?);
wenzelm
parents: 41413
diff changeset
     3
  Hotel_Example
c15ef1b85035 modernized imports (untested!?);
wenzelm
parents: 41413
diff changeset
     4
  "~~/src/HOL/Library/Predicate_Compile_Alternative_Defs"
c15ef1b85035 modernized imports (untested!?);
wenzelm
parents: 41413
diff changeset
     5
  "~~/src/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
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
    14
lemma [code_pred_inline]: "(op -) == (%A B x. A x \<and> \<not> B x)"
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"
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
    34
  "cardsp (e # s) g k = (let C = cardsp s g in case e of Check_in g' r c => if g' = g then k = c \<or> C k else C k | _ => C k)"
46905
6b1c0a80a57a prefer abs_def over def_raw;
wenzelm
parents: 45970
diff changeset
    35
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
    36
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
    37
definition
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
    38
  "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
    39
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
    40
lemma [code_pred_def]:
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
    41
  "isinp [] r g = False"
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
    42
  "isinp (e # s) r g =
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
    43
  (let G = isinp s r
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
    44
   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
    45
    | 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
    46
    | Exit g' r' => if r' = r then ((g \<noteq> g') \<and> G g) else G g)"
46905
6b1c0a80a57a prefer abs_def over def_raw;
wenzelm
parents: 45970
diff changeset
    47
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
    48
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
    49
declare hotel.simps(1)[code_pred_def]
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
    50
lemma [code_pred_def]:
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
    51
"hotel (e # s) =
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
    52
  (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
    53
  | 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
    54
  | Exit g r => isinp s r g))"
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
    55
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
    56
by (auto split: event.split)
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
    57
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
    58
declare List.member_rec[code_pred_def]
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
lemma [code_pred_def]: "no_Check_in s r = (~ (EX g c. List.member s (Check_in g r c)))"
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
    61
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
    62
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
    63
lemma [code_pred_def]: "feels_safe s r =
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
    64
(EX s\<^isub>1 s\<^isub>2 s\<^isub>3 g c c'.
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
    65
    s =
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
    66
    s\<^isub>3 @
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
    67
    [Enter g r c] @ s\<^isub>2 @ [Check_in g r c'] @ s\<^isub>1 &
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
    68
    no_Check_in (s\<^isub>3 @ s\<^isub>2) r &
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
    69
    (\<not> (\<exists> g'. isinp (s\<^isub>2 @ [Check_in g r c] @ s\<^isub>1) r g')))"
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
    70
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
    71
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
    72
setup {* 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
    73
  {ensure_groundness = true,
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
    74
  limit_globally = NONE,
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
    75
  limited_types = [],
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
    76
  limited_predicates = [],
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
    77
  replacing = [],
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
    78
  manual_reorder = []}) *}
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
    79
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
    80
values 40 "{s. hotel s}"
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
    81
43885
7caa1139b4e5 adapting prolog-based tester
bulwahn
parents: 41956
diff changeset
    82
setup {* Context.theory_map (Quickcheck.add_tester ("prolog", (Code_Prolog.active, Code_Prolog.test_goals))) *}
40104
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
    83
45970
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
    84
lemma "\<lbrakk> hotel s; isinp s r g \<rbrakk> \<Longrightarrow> owns s r = Some g"
40924
a9be7f26b4e6 adapting predicate_compile_quickcheck
bulwahn
parents: 40137
diff changeset
    85
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
    86
oops
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
    87
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
    88
section {* Manual setup to find the counterexample *}
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
    89
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
    90
setup {* 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
    91
  {ensure_groundness = true,
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
    92
   limit_globally = NONE,
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
    93
   limited_types = [],
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
    94
   limited_predicates = [(["hotel"], 4)],
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
    95
   replacing = [(("hotel", "limited_hotel"), "quickcheck")],
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
    96
   manual_reorder = []}) *}
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
    97
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
    98
lemma
45970
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
    99
  "hotel s ==> feels_safe s r ==> isinp s r g ==> owns s r = Some g"
40924
a9be7f26b4e6 adapting predicate_compile_quickcheck
bulwahn
parents: 40137
diff changeset
   100
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
   101
oops
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
setup {* 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
   104
  {ensure_groundness = true,
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
   105
   limit_globally = NONE,
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
   106
   limited_types = [],
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
   107
   limited_predicates = [(["hotel"], 5)],
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
   108
   replacing = [(("hotel", "limited_hotel"), "quickcheck")],
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
   109
   manual_reorder = []}) *}
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
   110
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
   111
lemma
45970
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
   112
  "hotel s ==> feels_safe s r ==> isinp s r g ==> owns s r = Some g"
40924
a9be7f26b4e6 adapting predicate_compile_quickcheck
bulwahn
parents: 40137
diff changeset
   113
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
   114
oops
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
section {* Using a global limit for limiting the execution *} 
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
   117
45970
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
   118
text {* A global depth limit of 10 does not suffice to find the counterexample. *}
40104
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
   119
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
   120
setup {*
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
   121
  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
   122
  {ensure_groundness = true,
45970
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
   123
  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
   124
  limited_types = [],
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
   125
  limited_predicates = [],
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
   126
  replacing = [],
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
   127
  manual_reorder = []})
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
   128
*}
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
   129
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
   130
lemma
45970
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
   131
  "hotel s ==> feels_safe s r ==> isinp s r g ==> owns s r = Some g"
40924
a9be7f26b4e6 adapting predicate_compile_quickcheck
bulwahn
parents: 40137
diff changeset
   132
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
   133
oops
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
   134
45970
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
   135
text {* But a global depth limit of 11 does. *}
40104
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
   136
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
   137
setup {*
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
   138
  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
   139
  {ensure_groundness = true,
45970
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
   140
  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
   141
  limited_types = [],
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
   142
  limited_predicates = [],
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
   143
  replacing = [],
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
   144
  manual_reorder = []})
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
   145
*}
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
   146
82873a6f2b81 splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
diff changeset
   147
lemma
45970
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 43885
diff changeset
   148
  "hotel s ==> feels_safe s r ==> isinp s r g ==> owns s r = Some g"
40924
a9be7f26b4e6 adapting predicate_compile_quickcheck
bulwahn
parents: 40137
diff changeset
   149
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
   150
oops
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
end