src/HOL/Quickcheck_Examples/Hotel_Example.thy
author wenzelm
Thu, 01 Sep 2016 16:13:46 +0200
changeset 63752 79f11158dcc4
parent 63177 6c05c4632949
child 66453 cc19f7ca2ed6
permissions -rw-r--r--
merged
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
48222
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
     1
theory Hotel_Example
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
     2
imports Main "~~/src/HOL/Library/Predicate_Compile_Quickcheck"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
     3
begin
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
     4
58310
91ea607a34d8 updated news
blanchet
parents: 58148
diff changeset
     5
datatype guest = Guest0 | Guest1
91ea607a34d8 updated news
blanchet
parents: 58148
diff changeset
     6
datatype key = Key0 | Key1 | Key2 | Key3
91ea607a34d8 updated news
blanchet
parents: 58148
diff changeset
     7
datatype room = Room0
48222
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
     8
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
     9
type_synonym card = "key * key"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    10
58310
91ea607a34d8 updated news
blanchet
parents: 58148
diff changeset
    11
datatype event =
58148
9764b994a421 use 'datatype_new' in Quickcheck examples
blanchet
parents: 57544
diff changeset
    12
  Check_in guest room card
9764b994a421 use 'datatype_new' in Quickcheck examples
blanchet
parents: 57544
diff changeset
    13
| Enter guest room card
9764b994a421 use 'datatype_new' in Quickcheck examples
blanchet
parents: 57544
diff changeset
    14
| Exit guest room
48222
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    15
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    16
definition initk :: "room \<Rightarrow> key"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    17
  where "initk = (%r. Key0)"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    18
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    19
declare initk_def[code_pred_def, code]
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    20
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    21
primrec owns :: "event list \<Rightarrow> room \<Rightarrow> guest option"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    22
where
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    23
  "owns [] r = None"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    24
| "owns (e#s) r = (case e of
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    25
    Check_in g r' c \<Rightarrow> if r' = r then Some g else owns s r |
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    26
    Enter g r' c \<Rightarrow> owns s r |
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    27
    Exit g r' \<Rightarrow> owns s r)"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    28
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    29
primrec currk :: "event list \<Rightarrow> room \<Rightarrow> key"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    30
where
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    31
  "currk [] r = initk r"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    32
| "currk (e#s) r = (let k = currk s r in
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    33
    case e of Check_in g r' (k1, k2) \<Rightarrow> if r' = r then k2 else k
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    34
            | Enter g r' c \<Rightarrow> k
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    35
            | Exit g r \<Rightarrow> k)"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    36
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    37
primrec issued :: "event list \<Rightarrow> key set"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    38
where
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    39
  "issued [] = range initk"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    40
| "issued (e#s) = issued s \<union>
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    41
  (case e of Check_in g r (k1, k2) \<Rightarrow> {k2} | Enter g r c \<Rightarrow> {} | Exit g r \<Rightarrow> {})"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    42
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    43
primrec cards :: "event list \<Rightarrow> guest \<Rightarrow> card set"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    44
where
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    45
  "cards [] g = {}"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    46
| "cards (e#s) g = (let C = cards s g in
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    47
                    case e of Check_in g' r c \<Rightarrow> if g' = g then insert c C
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    48
                                                else C
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    49
                            | Enter g r c \<Rightarrow> C
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    50
                            | Exit g r \<Rightarrow> C)"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    51
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    52
primrec roomk :: "event list \<Rightarrow> room \<Rightarrow> key"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    53
where
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    54
  "roomk [] r = initk r"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    55
| "roomk (e#s) r = (let k = roomk s r in
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    56
    case e of Check_in g r' c \<Rightarrow> k
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    57
            | Enter g r' (x,y) \<Rightarrow> if r' = r (*& x = k*) then y else k
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    58
            | Exit g r \<Rightarrow> k)"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    59
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    60
primrec isin :: "event list \<Rightarrow> room \<Rightarrow> guest set"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    61
where
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    62
  "isin [] r = {}"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    63
| "isin (e#s) r = (let G = isin s r in
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    64
                 case e of Check_in g r c \<Rightarrow> G
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    65
                 | Enter g r' c \<Rightarrow> if r' = r then {g} \<union> G else G
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    66
                 | Exit g r' \<Rightarrow> if r'=r then G - {g} else G)"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    67
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    68
primrec hotel :: "event list \<Rightarrow> bool"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    69
where
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    70
  "hotel []  = True"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    71
| "hotel (e # s) = (hotel s & (case e of
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    72
  Check_in g r (k,k') \<Rightarrow> k = currk s r \<and> k' \<notin> issued s |
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    73
  Enter g r (k,k') \<Rightarrow> (k,k') : cards s g & (roomk s r : {k, k'}) |
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    74
  Exit g r \<Rightarrow> g : isin s r))"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    75
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    76
definition no_Check_in :: "event list \<Rightarrow> room \<Rightarrow> bool" where(*>*)
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    77
[code del]: "no_Check_in s r \<equiv> \<not>(\<exists>g c. Check_in g r c \<in> set s)"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    78
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    79
definition feels_safe :: "event list \<Rightarrow> room \<Rightarrow> bool"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    80
where
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51272
diff changeset
    81
  "feels_safe s r = (\<exists>s\<^sub>1 s\<^sub>2 s\<^sub>3 g c c'.
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51272
diff changeset
    82
   s = s\<^sub>3 @ [Enter g r c] @ s\<^sub>2 @ [Check_in g r c'] @ s\<^sub>1 \<and>
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51272
diff changeset
    83
   no_Check_in (s\<^sub>3 @ s\<^sub>2) r \<and> isin (s\<^sub>2 @ [Check_in g r c] @ s\<^sub>1) r = {})"
48222
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    84
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    85
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58310
diff changeset
    86
section \<open>Some setup\<close>
48222
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    87
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    88
lemma issued_nil: "issued [] = {Key0}"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    89
by (auto simp add: initk_def)
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    90
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    91
lemmas issued_simps[code] = issued_nil issued.simps(2)
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    92
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    93
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58310
diff changeset
    94
setup \<open>Predicate_Compile_Data.ignore_consts [@{const_name Set.member},
48222
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    95
  @{const_name "issued"}, @{const_name "cards"}, @{const_name "isin"},
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58310
diff changeset
    96
  @{const_name Collect}, @{const_name insert}]\<close>
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58310
diff changeset
    97
ML_val \<open>Core_Data.force_modes_and_compilations\<close>
48222
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    98
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
    99
fun find_first :: "('a => 'b option) => 'a list => 'b option"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   100
where
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   101
  "find_first f [] = None"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   102
| "find_first f (x # xs) = (case f x of Some y => Some y | None => find_first f xs)"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   103
57544
8840fa17e17c reactivate session Quickcheck_Examples
Andreas Lochbihler
parents: 53015
diff changeset
   104
axiomatization cps_of_set :: "'a set => ('a => term list option) => term list option"
8840fa17e17c reactivate session Quickcheck_Examples
Andreas Lochbihler
parents: 53015
diff changeset
   105
where cps_of_set_code [code]: "cps_of_set (set xs) f = find_first f xs"
48222
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   106
57544
8840fa17e17c reactivate session Quickcheck_Examples
Andreas Lochbihler
parents: 53015
diff changeset
   107
axiomatization pos_cps_of_set :: "'a set => ('a => (bool * term list) option) => natural => (bool * term list) option"
8840fa17e17c reactivate session Quickcheck_Examples
Andreas Lochbihler
parents: 53015
diff changeset
   108
where pos_cps_of_set_code [code]: "pos_cps_of_set (set xs) f i = find_first f xs"
48222
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   109
57544
8840fa17e17c reactivate session Quickcheck_Examples
Andreas Lochbihler
parents: 53015
diff changeset
   110
axiomatization find_first' :: "('b Quickcheck_Exhaustive.unknown => 'a Quickcheck_Exhaustive.three_valued)
48222
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   111
    => 'b list => 'a Quickcheck_Exhaustive.three_valued"
63177
6c05c4632949 clarified axiomatization;
wenzelm
parents: 63167
diff changeset
   112
where find_first'_Nil: "find_first' f [] = Quickcheck_Exhaustive.No_value"
6c05c4632949 clarified axiomatization;
wenzelm
parents: 63167
diff changeset
   113
  and find_first'_Cons: "find_first' f (x # xs) =
6c05c4632949 clarified axiomatization;
wenzelm
parents: 63167
diff changeset
   114
    (case f (Quickcheck_Exhaustive.Known x) of
6c05c4632949 clarified axiomatization;
wenzelm
parents: 63167
diff changeset
   115
      Quickcheck_Exhaustive.No_value => find_first' f xs
6c05c4632949 clarified axiomatization;
wenzelm
parents: 63167
diff changeset
   116
    | Quickcheck_Exhaustive.Value x => Quickcheck_Exhaustive.Value x
6c05c4632949 clarified axiomatization;
wenzelm
parents: 63167
diff changeset
   117
    | Quickcheck_Exhaustive.Unknown_value =>
6c05c4632949 clarified axiomatization;
wenzelm
parents: 63167
diff changeset
   118
        (case find_first' f xs of Quickcheck_Exhaustive.Value x =>
6c05c4632949 clarified axiomatization;
wenzelm
parents: 63167
diff changeset
   119
          Quickcheck_Exhaustive.Value x
6c05c4632949 clarified axiomatization;
wenzelm
parents: 63167
diff changeset
   120
        | _ => Quickcheck_Exhaustive.Unknown_value))"
6c05c4632949 clarified axiomatization;
wenzelm
parents: 63167
diff changeset
   121
6c05c4632949 clarified axiomatization;
wenzelm
parents: 63167
diff changeset
   122
lemmas find_first'_code [code] = find_first'_Nil find_first'_Cons
48222
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   123
57544
8840fa17e17c reactivate session Quickcheck_Examples
Andreas Lochbihler
parents: 53015
diff changeset
   124
axiomatization neg_cps_of_set :: "'a set => ('a Quickcheck_Exhaustive.unknown => term list Quickcheck_Exhaustive.three_valued) => natural => term list Quickcheck_Exhaustive.three_valued"
8840fa17e17c reactivate session Quickcheck_Examples
Andreas Lochbihler
parents: 53015
diff changeset
   125
where neg_cps_of_set_code [code]: "neg_cps_of_set (set xs) f i = find_first' f xs"
48222
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   126
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58310
diff changeset
   127
setup \<open>
48222
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   128
let
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   129
  val Fun = Predicate_Compile_Aux.Fun
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   130
  val Input = Predicate_Compile_Aux.Input
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   131
  val Output = Predicate_Compile_Aux.Output
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   132
  val Bool = Predicate_Compile_Aux.Bool
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   133
  val oi = Fun (Output, Fun (Input, Bool))
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   134
  val ii = Fun (Input, Fun (Input, Bool))
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   135
  fun of_set compfuns (Type ("fun", [T, _])) =
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   136
    case body_type (Predicate_Compile_Aux.mk_monadT compfuns T) of
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   137
      Type ("Quickcheck_Exhaustive.three_valued", _) => 
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   138
        Const(@{const_name neg_cps_of_set}, HOLogic.mk_setT T --> (Predicate_Compile_Aux.mk_monadT compfuns T))
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   139
    | _ => Const(@{const_name pos_cps_of_set}, HOLogic.mk_setT T --> (Predicate_Compile_Aux.mk_monadT compfuns T))
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   140
  fun member compfuns (U as Type ("fun", [T, _])) =
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   141
    (absdummy T (absdummy (HOLogic.mk_setT T) (Predicate_Compile_Aux.mk_if compfuns
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   142
      (Const (@{const_name "Set.member"}, T --> HOLogic.mk_setT T --> @{typ bool}) $ Bound 1 $ Bound 0))))
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   143
 
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   144
in
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   145
  Core_Data.force_modes_and_compilations @{const_name Set.member}
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   146
    [(oi, (of_set, false)), (ii, (member, false))]
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   147
end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58310
diff changeset
   148
\<close>
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58310
diff changeset
   149
section \<open>Property\<close>
48222
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   150
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   151
lemma "\<lbrakk> hotel s; g \<in> isin s r \<rbrakk> \<Longrightarrow> owns s r = Some g"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   152
quickcheck[tester = exhaustive, size = 6, expect = counterexample]
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   153
quickcheck[tester = smart_exhaustive, depth = 6, expect = counterexample]
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   154
oops
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   155
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   156
lemma
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   157
  "hotel s ==> feels_safe s r ==> g \<in> isin s r ==> owns s r = Some g"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   158
quickcheck[smart_exhaustive, depth = 10, allow_function_inversion, expect = counterexample]
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   159
oops
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   160
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58310
diff changeset
   161
section \<open>Refinement\<close>
48222
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   162
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   163
fun split_list
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   164
where
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   165
  "split_list [] = [([], [])]"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   166
| "split_list (z # zs) = (([], z # zs) # [(z # xs', ys'). (xs', ys') <- split_list zs])"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   167
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   168
lemma split_list: "((xs, ys) \<in> set (split_list zs)) = (zs = xs @ ys)"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   169
apply (induct zs arbitrary: xs ys)
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   170
apply fastforce
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   171
apply (case_tac xs)
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   172
apply auto
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   173
done
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   174
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   175
lemma [code]: "no_Check_in s r = list_all (%x. case x of Check_in g r' c => r \<noteq> r' | _ => True) s"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   176
unfolding no_Check_in_def list_all_iff
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   177
apply auto
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   178
apply (case_tac x)
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   179
apply auto
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   180
done
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   181
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   182
lemma [code]: "feels_safe s r = list_ex (%(s3, s2, s1, g, c, c'). no_Check_in (s3 @ s2) r &
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   183
    isin (s2 @ [Check_in g r c] @ s1) r = {}) ([(s3, s2, s1, g, c, c'). (s3, Enter g' r' c # r3) <- split_list s, r' = r, (s2, Check_in g r'' c' # s1) <- split_list r3, r'' = r, g = g'])"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   184
unfolding feels_safe_def list_ex_iff
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   185
by auto (metis split_list)+
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   186
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   187
lemma
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   188
  "hotel s ==> feels_safe s r ==> g \<in> isin s r ==> owns s r = Some g"
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   189
(* quickcheck[exhaustive, size = 9, timeout = 2000] -- maybe possible with a lot of time *)
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   190
quickcheck[narrowing, size = 7, expect = counterexample]
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   191
oops
fcca68383808 adding the hotel key card example in Quickcheck-Examples
bulwahn
parents:
diff changeset
   192
48255
968602739b54 fixed typo
bulwahn
parents: 48222
diff changeset
   193
end