| author | nipkow | 
| Mon, 13 Sep 2010 11:13:15 +0200 | |
| changeset 39302 | d7728f65b353 | 
| parent 39200 | bb93713b0925 | 
| child 39463 | 7ce0ed8dc4d6 | 
| permissions | -rw-r--r-- | 
| 38730 | 1 | theory Hotel_Example | 
| 2 | imports Predicate_Compile_Alternative_Defs Code_Prolog | |
| 3 | begin | |
| 4 | ||
| 5 | datatype guest = Guest0 | Guest1 | |
| 6 | datatype key = Key0 | Key1 | Key2 | Key3 | |
| 7 | datatype room = Room0 | |
| 8 | ||
| 9 | types card = "key * key" | |
| 10 | ||
| 11 | datatype event = | |
| 12 | Check_in guest room card | Enter guest room card | Exit guest room | |
| 13 | ||
| 14 | definition initk :: "room \<Rightarrow> key" | |
| 15 | where "initk = (%r. Key0)" | |
| 16 | ||
| 17 | declare initk_def[code_pred_def, code] | |
| 18 | ||
| 19 | primrec owns :: "event list \<Rightarrow> room \<Rightarrow> guest option" | |
| 20 | where | |
| 21 | "owns [] r = None" | |
| 22 | | "owns (e#s) r = (case e of | |
| 23 | Check_in g r' c \<Rightarrow> if r' = r then Some g else owns s r | | |
| 24 | Enter g r' c \<Rightarrow> owns s r | | |
| 25 | Exit g r' \<Rightarrow> owns s r)" | |
| 26 | ||
| 27 | primrec currk :: "event list \<Rightarrow> room \<Rightarrow> key" | |
| 28 | where | |
| 29 | "currk [] r = initk r" | |
| 30 | | "currk (e#s) r = (let k = currk s r in | |
| 31 | case e of Check_in g r' (k1, k2) \<Rightarrow> if r' = r then k2 else k | |
| 32 | | Enter g r' c \<Rightarrow> k | |
| 33 | | Exit g r \<Rightarrow> k)" | |
| 34 | ||
| 35 | primrec issued :: "event list \<Rightarrow> key set" | |
| 36 | where | |
| 37 | "issued [] = range initk" | |
| 38 | | "issued (e#s) = issued s \<union> | |
| 39 |   (case e of Check_in g r (k1, k2) \<Rightarrow> {k2} | Enter g r c \<Rightarrow> {} | Exit g r \<Rightarrow> {})"
 | |
| 40 | ||
| 41 | primrec cards :: "event list \<Rightarrow> guest \<Rightarrow> card set" | |
| 42 | where | |
| 43 |   "cards [] g = {}"
 | |
| 44 | | "cards (e#s) g = (let C = cards s g in | |
| 45 | case e of Check_in g' r c \<Rightarrow> if g' = g then insert c C | |
| 46 | else C | |
| 47 | | Enter g r c \<Rightarrow> C | |
| 48 | | Exit g r \<Rightarrow> C)" | |
| 49 | ||
| 50 | primrec roomk :: "event list \<Rightarrow> room \<Rightarrow> key" | |
| 51 | where | |
| 52 | "roomk [] r = initk r" | |
| 53 | | "roomk (e#s) r = (let k = roomk s r in | |
| 54 | case e of Check_in g r' c \<Rightarrow> k | |
| 55 | | Enter g r' (x,y) \<Rightarrow> if r' = r (*& x = k*) then y else k | |
| 56 | | Exit g r \<Rightarrow> k)" | |
| 57 | ||
| 58 | primrec isin :: "event list \<Rightarrow> room \<Rightarrow> guest set" | |
| 59 | where | |
| 60 |   "isin [] r = {}"
 | |
| 61 | | "isin (e#s) r = (let G = isin s r in | |
| 62 | case e of Check_in g r c \<Rightarrow> G | |
| 63 |                  | Enter g r' c \<Rightarrow> if r' = r then {g} \<union> G else G
 | |
| 64 |                  | Exit g r' \<Rightarrow> if r'=r then G - {g} else G)"
 | |
| 65 | ||
| 66 | primrec hotel :: "event list \<Rightarrow> bool" | |
| 67 | where | |
| 68 | "hotel [] = True" | |
| 69 | | "hotel (e # s) = (hotel s & (case e of | |
| 70 | Check_in g r (k,k') \<Rightarrow> k = currk s r \<and> k' \<notin> issued s | | |
| 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: 
38733diff
changeset | 72 | Exit g r \<Rightarrow> g : isin s r))" | 
| 38730 | 73 | |
| 74 | lemma issued_nil: "issued [] = {Key0}"
 | |
| 75 | by (auto simp add: initk_def) | |
| 76 | ||
| 77 | lemmas issued_simps[code, code_pred_def] = issued_nil issued.simps(2) | |
| 78 | ||
| 79 | declare Let_def[code_pred_inline] | |
| 80 | ||
| 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: 
39200diff
changeset | 82 | by (auto simp add: insert_iff[unfolded mem_def] fun_eq_iff intro!: eq_reflection) | 
| 38730 | 83 | |
| 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: 
39200diff
changeset | 85 | by (auto simp add: Diff_iff[unfolded mem_def] fun_eq_iff intro!: eq_reflection) | 
| 38730 | 86 | |
| 38950 
62578950e748
storing options for prolog code generation in the theory
 bulwahn parents: 
38949diff
changeset | 87 | setup {* Code_Prolog.map_code_options (K
 | 
| 38949 | 88 |   {ensure_groundness = true,
 | 
| 89 | limited_types = [], | |
| 90 | limited_predicates = [], | |
| 91 | replacing = [], | |
| 38963 | 92 | manual_reorder = [], | 
| 39189 | 93 | timeout = Time.fromSeconds 10, | 
| 38950 
62578950e748
storing options for prolog code generation in the theory
 bulwahn parents: 
38949diff
changeset | 94 | prolog_system = Code_Prolog.SWI_PROLOG}) *} | 
| 38730 | 95 | |
| 38734 
e5508a74b11f
changing hotel trace definition; adding simple handling of numerals on natural numbers
 bulwahn parents: 
38733diff
changeset | 96 | values 40 "{s. hotel s}"
 | 
| 38730 | 97 | |
| 38731 
2c8a595af43e
invocation of values for prolog execution does not require invocation of code_pred anymore
 bulwahn parents: 
38730diff
changeset | 98 | |
| 38733 
4b8fd91ea59a
added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
 bulwahn parents: 
38731diff
changeset | 99 | setup {* Quickcheck.add_generator ("prolog", Code_Prolog.quickcheck) *}
 | 
| 
4b8fd91ea59a
added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
 bulwahn parents: 
38731diff
changeset | 100 | |
| 
4b8fd91ea59a
added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
 bulwahn parents: 
38731diff
changeset | 101 | lemma "\<lbrakk> hotel s; g \<in> isin s r \<rbrakk> \<Longrightarrow> owns s r = Some g" | 
| 
4b8fd91ea59a
added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
 bulwahn parents: 
38731diff
changeset | 102 | quickcheck[generator = code, iterations = 100000, report] | 
| 38949 | 103 | 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: 
38731diff
changeset | 104 | oops | 
| 
4b8fd91ea59a
added quickcheck generator for prolog generation; first example of counterexample search with prolog for hotel key card system
 bulwahn parents: 
38731diff
changeset | 105 | |
| 38731 
2c8a595af43e
invocation of values for prolog execution does not require invocation of code_pred anymore
 bulwahn parents: 
38730diff
changeset | 106 | |
| 38953 
0c38eb5fc4ca
added further hotel key card attack in example file
 bulwahn parents: 
38950diff
changeset | 107 | definition no_Check_in :: "event list \<Rightarrow> room \<Rightarrow> bool" where(*>*) | 
| 
0c38eb5fc4ca
added further hotel key card attack in example file
 bulwahn parents: 
38950diff
changeset | 108 | [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: 
38950diff
changeset | 109 | |
| 
0c38eb5fc4ca
added further hotel key card attack in example file
 bulwahn parents: 
38950diff
changeset | 110 | |
| 
0c38eb5fc4ca
added further hotel key card attack in example file
 bulwahn parents: 
38950diff
changeset | 111 | definition feels_safe :: "event list \<Rightarrow> room \<Rightarrow> bool" | 
| 
0c38eb5fc4ca
added further hotel key card attack in example file
 bulwahn parents: 
38950diff
changeset | 112 | where | 
| 
0c38eb5fc4ca
added further hotel key card attack in example file
 bulwahn parents: 
38950diff
changeset | 113 | "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: 
38950diff
changeset | 114 | 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: 
38950diff
changeset | 115 |    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: 
38950diff
changeset | 116 | |
| 
0c38eb5fc4ca
added further hotel key card attack in example file
 bulwahn parents: 
38950diff
changeset | 117 | setup {* Code_Prolog.map_code_options (K 
 | 
| 
0c38eb5fc4ca
added further hotel key card attack in example file
 bulwahn parents: 
38950diff
changeset | 118 |   {ensure_groundness = true,
 | 
| 
0c38eb5fc4ca
added further hotel key card attack in example file
 bulwahn parents: 
38950diff
changeset | 119 | limited_types = [], | 
| 38963 | 120 | limited_predicates = [(["hotel"], 5)], | 
| 38953 
0c38eb5fc4ca
added further hotel key card attack in example file
 bulwahn parents: 
38950diff
changeset | 121 |    replacing = [(("hotel", "limited_hotel"), "quickcheck")],
 | 
| 38963 | 122 | manual_reorder = [], | 
| 39189 | 123 | timeout = Time.fromSeconds 10, | 
| 38953 
0c38eb5fc4ca
added further hotel key card attack in example file
 bulwahn parents: 
38950diff
changeset | 124 | prolog_system = Code_Prolog.SWI_PROLOG}) *} | 
| 
0c38eb5fc4ca
added further hotel key card attack in example file
 bulwahn parents: 
38950diff
changeset | 125 | |
| 
0c38eb5fc4ca
added further hotel key card attack in example file
 bulwahn parents: 
38950diff
changeset | 126 | 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: 
38953diff
changeset | 127 | "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: 
38953diff
changeset | 128 | quickcheck[generator = prolog, iterations = 1, expect = counterexample] | 
| 38953 
0c38eb5fc4ca
added further hotel key card attack in example file
 bulwahn parents: 
38950diff
changeset | 129 | oops | 
| 
0c38eb5fc4ca
added further hotel key card attack in example file
 bulwahn parents: 
38950diff
changeset | 130 | |
| 38730 | 131 | end |