author | paulson <lp15@cam.ac.uk> |
Thu, 27 Feb 2014 16:07:21 +0000 | |
changeset 55775 | 1557a391a858 |
parent 53015 | a1119cf551e8 |
child 58249 | 180f1b3508ed |
permissions | -rw-r--r-- |
38730 | 1 |
theory Hotel_Example |
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
39799
diff
changeset
|
2 |
imports Main |
38730 | 3 |
begin |
4 |
||
5 |
datatype guest = Guest0 | Guest1 |
|
6 |
datatype key = Key0 | Key1 | Key2 | Key3 |
|
7 |
datatype room = Room0 |
|
8 |
||
42463 | 9 |
type_synonym card = "key * key" |
38730 | 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:
38733
diff
changeset
|
72 |
Exit g r \<Rightarrow> g : isin s r))" |
38730 | 73 |
|
38953
0c38eb5fc4ca
added further hotel key card attack in example file
bulwahn
parents:
38950
diff
changeset
|
74 |
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
|
75 |
[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
|
76 |
|
0c38eb5fc4ca
added further hotel key card attack in example file
bulwahn
parents:
38950
diff
changeset
|
77 |
definition feels_safe :: "event list \<Rightarrow> room \<Rightarrow> bool" |
0c38eb5fc4ca
added further hotel key card attack in example file
bulwahn
parents:
38950
diff
changeset
|
78 |
where |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
42463
diff
changeset
|
79 |
"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:
42463
diff
changeset
|
80 |
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:
42463
diff
changeset
|
81 |
no_Check_in (s\<^sub>3 @ s\<^sub>2) r \<and> isin (s\<^sub>2 @ [Check_in g r c] @ s\<^sub>1) r = {})" |
38953
0c38eb5fc4ca
added further hotel key card attack in example file
bulwahn
parents:
38950
diff
changeset
|
82 |
|
39799
fdbea66eae4b
finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents:
39463
diff
changeset
|
83 |
|
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
39799
diff
changeset
|
84 |
section {* Some setup *} |
39799
fdbea66eae4b
finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents:
39463
diff
changeset
|
85 |
|
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
39799
diff
changeset
|
86 |
lemma issued_nil: "issued [] = {Key0}" |
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
39799
diff
changeset
|
87 |
by (auto simp add: initk_def) |
39799
fdbea66eae4b
finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents:
39463
diff
changeset
|
88 |
|
40104
82873a6f2b81
splitting Hotel Key card example into specification and the two tests for counter example generation
bulwahn
parents:
39799
diff
changeset
|
89 |
lemmas issued_simps[code, code_pred_def] = issued_nil issued.simps(2) |
39799
fdbea66eae4b
finding the counterexample with different options (manually limited precisely, manually limited global, automatically limited global)
bulwahn
parents:
39463
diff
changeset
|
90 |
|
38730 | 91 |
end |