src/HOL/Predicate_Compile_Examples/Hotel_Example.thy
author haftmann
Mon, 04 Oct 2010 14:46:49 +0200
changeset 39919 9f6503aaa77d
parent 39799 fdbea66eae4b
child 40104 82873a6f2b81
permissions -rw-r--r--
adjusted to inductive characterization of sorted
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
38730
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
     1
theory Hotel_Example
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
     2
imports Predicate_Compile_Alternative_Defs Code_Prolog
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
     3
begin
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
     4
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
     5
datatype guest = Guest0 | Guest1
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
     6
datatype key = Key0 | Key1 | Key2 | Key3
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
     7
datatype room = Room0
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
     8
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
     9
types card = "key * key"
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    10
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    11
datatype event =
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    12
   Check_in guest room card | Enter guest room card | Exit guest room
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    13
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    14
definition initk :: "room \<Rightarrow> key"
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    15
  where "initk = (%r. Key0)"
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    16
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    17
declare initk_def[code_pred_def, code]
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    18
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    19
primrec owns :: "event list \<Rightarrow> room \<Rightarrow> guest option"
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    20
where
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    21
  "owns [] r = None"
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    22
| "owns (e#s) r = (case e of
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    23
    Check_in g r' c \<Rightarrow> if r' = r then Some g else owns s r |
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    24
    Enter g r' c \<Rightarrow> owns s r |
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    25
    Exit g r' \<Rightarrow> owns s r)"
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    26
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    27
primrec currk :: "event list \<Rightarrow> room \<Rightarrow> key"
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    28
where
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    29
  "currk [] r = initk r"
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    30
| "currk (e#s) r = (let k = currk s r in
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    31
    case e of Check_in g r' (k1, k2) \<Rightarrow> if r' = r then k2 else k
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    32
            | Enter g r' c \<Rightarrow> k
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    33
            | Exit g r \<Rightarrow> k)"
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    34
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    35
primrec issued :: "event list \<Rightarrow> key set"
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    36
where
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    37
  "issued [] = range initk"
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    38
| "issued (e#s) = issued s \<union>
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    39
  (case e of Check_in g r (k1, k2) \<Rightarrow> {k2} | Enter g r c \<Rightarrow> {} | Exit g r \<Rightarrow> {})"
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    40
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    41
primrec cards :: "event list \<Rightarrow> guest \<Rightarrow> card set"
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    42
where
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    43
  "cards [] g = {}"
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    44
| "cards (e#s) g = (let C = cards s g in
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    45
                    case e of Check_in g' r c \<Rightarrow> if g' = g then insert c C
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    46
                                                else C
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    47
                            | Enter g r c \<Rightarrow> C
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    48
                            | Exit g r \<Rightarrow> C)"
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    49
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    50
primrec roomk :: "event list \<Rightarrow> room \<Rightarrow> key"
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    51
where
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    52
  "roomk [] r = initk r"
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    53
| "roomk (e#s) r = (let k = roomk s r in
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    54
    case e of Check_in g r' c \<Rightarrow> k
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    55
            | Enter g r' (x,y) \<Rightarrow> if r' = r (*& x = k*) then y else k
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    56
            | Exit g r \<Rightarrow> k)"
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    57
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    58
primrec isin :: "event list \<Rightarrow> room \<Rightarrow> guest set"
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    59
where
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    60
  "isin [] r = {}"
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    61
| "isin (e#s) r = (let G = isin s r in
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    62
                 case e of Check_in g r c \<Rightarrow> G
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    63
                 | Enter g r' c \<Rightarrow> if r' = r then {g} \<union> G else G
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    64
                 | Exit g r' \<Rightarrow> if r'=r then G - {g} else G)"
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    65
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    66
primrec hotel :: "event list \<Rightarrow> bool"
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    67
where
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    68
  "hotel []  = True"
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    69
| "hotel (e # s) = (hotel s & (case e of
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    70
  Check_in g r (k,k') \<Rightarrow> k = currk s r \<and> k' \<notin> issued s |
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    71
  Enter g r (k,k') \<Rightarrow> (k,k') : cards s g & (roomk s r : {k, k'}) |
38734
e5508a74b11f changing hotel trace definition; adding simple handling of numerals on natural numbers
bulwahn
parents: 38733
diff changeset
    72
  Exit g r \<Rightarrow> g : isin s r))"
38730
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    73
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    74
lemma issued_nil: "issued [] = {Key0}"
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    75
by (auto simp add: initk_def)
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    76
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    77
lemmas issued_simps[code, code_pred_def] = issued_nil issued.simps(2)
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    78
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    79
declare Let_def[code_pred_inline]
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    80
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    81
lemma [code_pred_inline]: "insert == (%y A x. y = x | A x)"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39200
diff changeset
    82
by (auto simp add: insert_iff[unfolded mem_def] fun_eq_iff intro!: eq_reflection)
38730
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    83
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    84
lemma [code_pred_inline]: "(op -) == (%A B x. A x \<and> \<not> B x)"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39200
diff changeset
    85
by (auto simp add: Diff_iff[unfolded mem_def] fun_eq_iff intro!: eq_reflection)
38730
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    86
38950
62578950e748 storing options for prolog code generation in the theory
bulwahn
parents: 38949
diff changeset
    87
setup {* Code_Prolog.map_code_options (K
38949
1afa9e89c885 adapting example files to latest changes
bulwahn
parents: 38734
diff changeset
    88
  {ensure_groundness = true,
39799
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
    89
  limit_globally = NONE,
38949
1afa9e89c885 adapting example files to latest changes
bulwahn
parents: 38734
diff changeset
    90
  limited_types = [],
1afa9e89c885 adapting example files to latest changes
bulwahn
parents: 38734
diff changeset
    91
  limited_predicates = [],
1afa9e89c885 adapting example files to latest changes
bulwahn
parents: 38734
diff changeset
    92
  replacing = [],
39463
7ce0ed8dc4d6 adapting examples
bulwahn
parents: 39302
diff changeset
    93
  manual_reorder = []}) *}
38730
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    94
38734
e5508a74b11f changing hotel trace definition; adding simple handling of numerals on natural numbers
bulwahn
parents: 38733
diff changeset
    95
values 40 "{s. hotel s}"
38730
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
    96
38731
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38730
diff changeset
    97
39463
7ce0ed8dc4d6 adapting examples
bulwahn
parents: 39302
diff changeset
    98
setup {* Context.theory_map (Quickcheck.add_generator ("prolog", Code_Prolog.quickcheck)) *}
38733
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38731
diff changeset
    99
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38731
diff changeset
   100
lemma "\<lbrakk> hotel s; g \<in> isin s r \<rbrakk> \<Longrightarrow> owns s r = Some g"
39799
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   101
(*quickcheck[generator = code, iterations = 100000, report]*)
38949
1afa9e89c885 adapting example files to latest changes
bulwahn
parents: 38734
diff changeset
   102
quickcheck[generator = prolog, iterations = 1, expect = counterexample]
38733
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38731
diff changeset
   103
oops
4b8fd91ea59a added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
bulwahn
parents: 38731
diff changeset
   104
38731
2c8a595af43e invocation of values for prolog execution does not require invocation of code_pred anymore
bulwahn
parents: 38730
diff changeset
   105
38953
0c38eb5fc4ca added further hotel key card attack in example file
bulwahn
parents: 38950
diff changeset
   106
definition no_Check_in :: "event list \<Rightarrow> room \<Rightarrow> bool" where(*>*)
0c38eb5fc4ca added further hotel key card attack in example file
bulwahn
parents: 38950
diff changeset
   107
[code del]: "no_Check_in s r \<equiv> \<not>(\<exists>g c. Check_in g r c \<in> set s)"
0c38eb5fc4ca added further hotel key card attack in example file
bulwahn
parents: 38950
diff changeset
   108
0c38eb5fc4ca added further hotel key card attack in example file
bulwahn
parents: 38950
diff changeset
   109
0c38eb5fc4ca added further hotel key card attack in example file
bulwahn
parents: 38950
diff changeset
   110
definition feels_safe :: "event list \<Rightarrow> room \<Rightarrow> bool"
0c38eb5fc4ca added further hotel key card attack in example file
bulwahn
parents: 38950
diff changeset
   111
where
0c38eb5fc4ca added further hotel key card attack in example file
bulwahn
parents: 38950
diff changeset
   112
  "feels_safe s r = (\<exists>s\<^isub>1 s\<^isub>2 s\<^isub>3 g c c'.
0c38eb5fc4ca added further hotel key card attack in example file
bulwahn
parents: 38950
diff changeset
   113
   s = s\<^isub>3 @ [Enter g r c] @ s\<^isub>2 @ [Check_in g r c'] @ s\<^isub>1 \<and>
0c38eb5fc4ca added further hotel key card attack in example file
bulwahn
parents: 38950
diff changeset
   114
   no_Check_in (s\<^isub>3 @ s\<^isub>2) r \<and> isin (s\<^isub>2 @ [Check_in g r c] @ s\<^isub>1) r = {})"
0c38eb5fc4ca added further hotel key card attack in example file
bulwahn
parents: 38950
diff changeset
   115
39799
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   116
section {* Manual setup to find the counterexample *}
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   117
38953
0c38eb5fc4ca added further hotel key card attack in example file
bulwahn
parents: 38950
diff changeset
   118
setup {* Code_Prolog.map_code_options (K 
0c38eb5fc4ca added further hotel key card attack in example file
bulwahn
parents: 38950
diff changeset
   119
  {ensure_groundness = true,
39799
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   120
   limit_globally = NONE,
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   121
   limited_types = [],
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   122
   limited_predicates = [(["hotel"], 4)],
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   123
   replacing = [(("hotel", "limited_hotel"), "quickcheck")],
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   124
   manual_reorder = []}) *}
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   125
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   126
lemma
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   127
  "hotel s ==> feels_safe s r ==> g \<in> isin s r ==> owns s r = Some g"
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   128
quickcheck[generator = prolog, iterations = 1, expect = no_counterexample]
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   129
oops
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   130
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   131
setup {* Code_Prolog.map_code_options (K 
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   132
  {ensure_groundness = true,
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   133
   limit_globally = NONE,
38953
0c38eb5fc4ca added further hotel key card attack in example file
bulwahn
parents: 38950
diff changeset
   134
   limited_types = [],
38963
b5d126d7be4b adapting and tuning example theories
bulwahn
parents: 38954
diff changeset
   135
   limited_predicates = [(["hotel"], 5)],
38953
0c38eb5fc4ca added further hotel key card attack in example file
bulwahn
parents: 38950
diff changeset
   136
   replacing = [(("hotel", "limited_hotel"), "quickcheck")],
39463
7ce0ed8dc4d6 adapting examples
bulwahn
parents: 39302
diff changeset
   137
   manual_reorder = []}) *}
38953
0c38eb5fc4ca added further hotel key card attack in example file
bulwahn
parents: 38950
diff changeset
   138
0c38eb5fc4ca added further hotel key card attack in example file
bulwahn
parents: 38950
diff changeset
   139
lemma
38954
80ce658600b0 changing order of premises generated when flattening functions in premises; adapting example for second attack for hotel key card system
bulwahn
parents: 38953
diff changeset
   140
  "hotel s ==> feels_safe s r ==> g \<in> isin s r ==> owns s r = Some g"
80ce658600b0 changing order of premises generated when flattening functions in premises; adapting example for second attack for hotel key card system
bulwahn
parents: 38953
diff changeset
   141
quickcheck[generator = prolog, iterations = 1, expect = counterexample]
38953
0c38eb5fc4ca added further hotel key card attack in example file
bulwahn
parents: 38950
diff changeset
   142
oops
0c38eb5fc4ca added further hotel key card attack in example file
bulwahn
parents: 38950
diff changeset
   143
39799
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   144
section {* Simulating a global depth limit manually by limiting all predicates *}
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   145
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   146
setup {*
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   147
  Code_Prolog.map_code_options (K
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   148
  {ensure_groundness = true,
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   149
  limit_globally = NONE,
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   150
  limited_types = [],
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   151
  limited_predicates = [(["feelssafe", "feelssafeaux", "noCheckin", "noCheckinaux", "appendP", "ownsP", "hotel", "hotelaux", "hotelauxaux", "roomkP", "issued", "currkP", "initkP",
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   152
    "cards", "cardsauxauxaux", "cardsauxaux", "cardsaux", "isin", "isinauxauxa", "isinauxauxaux", "isinauxaux", "isinaux", "set"], 12)],
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   153
  replacing = [(("hotel", "limited_hotel"), "quickcheck"), (("feelssafe", "limited_feelssafe"), "quickcheck"), (("isin", "limited_isin"), "quickcheck")],
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   154
  manual_reorder = []})
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   155
*}
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   156
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   157
lemma
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   158
  "hotel s ==> feels_safe s r ==> g \<in> isin s r ==> owns s r = Some g"
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   159
quickcheck[generator = prolog, iterations = 1, expect = no_counterexample]
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   160
oops
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   161
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   162
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   163
setup {*
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   164
  Code_Prolog.map_code_options (K
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   165
  {ensure_groundness = true,
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   166
  limit_globally = NONE,
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   167
  limited_types = [],
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   168
  limited_predicates = [(["feelssafe", "feelssafeaux", "noCheckin", "noCheckinaux", "appendP", "ownsP", "hotel", "hotelaux", "hotelauxaux", "roomkP", "issued", "currkP", "initkP",
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   169
    "cards", "cardsauxauxaux", "cardsauxaux", "cardsaux", "isin", "isinauxauxa", "isinauxauxaux", "isinauxaux", "isinaux", "set"], 13)],
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   170
  replacing = [(("hotel", "limited_hotel"), "quickcheck"), (("feelssafe", "limited_feelssafe"), "quickcheck"), (("isin", "limited_isin"), "quickcheck")],
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   171
  manual_reorder = []})
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   172
*}
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   173
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   174
lemma
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   175
  "hotel s ==> feels_safe s r ==> g \<in> isin s r ==> owns s r = Some g"
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   176
quickcheck[generator = prolog, iterations = 1, expect = counterexample]
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   177
oops
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   178
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   179
section {* Using a global limit for limiting the execution *} 
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   180
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   181
text {* A global depth limit of 13 does not suffice to find the counterexample. *}
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   182
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   183
setup {*
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   184
  Code_Prolog.map_code_options (K
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   185
  {ensure_groundness = true,
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   186
  limit_globally = SOME 13,
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   187
  limited_types = [],
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   188
  limited_predicates = [],
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   189
  replacing = [],
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   190
  manual_reorder = []})
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   191
*}
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   192
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   193
lemma
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   194
  "hotel s ==> feels_safe s r ==> g \<in> isin s r ==> owns s r = Some g"
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   195
quickcheck[generator = prolog, iterations = 1, expect = no_counterexample]
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   196
oops
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   197
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   198
text {* But a global depth limit of 14 does. *}
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   199
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   200
setup {*
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   201
  Code_Prolog.map_code_options (K
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   202
  {ensure_groundness = true,
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   203
  limit_globally = SOME 14,
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   204
  limited_types = [],
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   205
  limited_predicates = [],
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   206
  replacing = [],
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   207
  manual_reorder = []})
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   208
*}
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   209
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   210
lemma
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   211
  "hotel s ==> feels_safe s r ==> g \<in> isin s r ==> owns s r = Some g"
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   212
quickcheck[generator = prolog, iterations = 1, expect = counterexample]
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   213
oops
fdbea66eae4b finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents: 39463
diff changeset
   214
38730
5bbdd9a9df62 adding hotel keycard example for prolog generation
bulwahn
parents:
diff changeset
   215
end