author | blanchet |
Tue, 03 Aug 2010 17:43:15 +0200 | |
changeset 38185 | b51677438b3a |
parent 35284 | 9edc2bd6d2bd |
child 40341 | 03156257040f |
permissions | -rw-r--r-- |
35076
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
1 |
(* Title: HOL/Nitpick_Examples/Hotel_Nits.thy |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
2 |
Author: Jasmin Blanchette, TU Muenchen |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
3 |
Copyright 2010 |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
4 |
|
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
5 |
Nitpick example based on Tobias Nipkow's hotel key card formalization. |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
6 |
*) |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
7 |
|
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
8 |
header {* Nitpick Example Based on Tobias Nipkow's Hotel Key Card |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
9 |
Formalization *} |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
10 |
|
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
11 |
theory Hotel_Nits |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
12 |
imports Main |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
13 |
begin |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
14 |
|
35284
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents:
35078
diff
changeset
|
15 |
nitpick_params [max_potential = 0, sat_solver = MiniSat_JNI, max_threads = 1, |
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents:
35078
diff
changeset
|
16 |
timeout = 120 s] |
35076
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
17 |
|
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
18 |
typedecl guest |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
19 |
typedecl key |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
20 |
typedecl room |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
21 |
|
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
22 |
types keycard = "key \<times> key" |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
23 |
|
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
24 |
record state = |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
25 |
owns :: "room \<Rightarrow> guest option" |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
26 |
currk :: "room \<Rightarrow> key" |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
27 |
issued :: "key set" |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
28 |
cards :: "guest \<Rightarrow> keycard set" |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
29 |
roomk :: "room \<Rightarrow> key" |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
30 |
isin :: "room \<Rightarrow> guest set" |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
31 |
safe :: "room \<Rightarrow> bool" |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
32 |
|
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
33 |
inductive_set reach :: "state set" where |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
34 |
init: |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
35 |
"inj initk \<Longrightarrow> |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
36 |
\<lparr>owns = (\<lambda>r. None), currk = initk, issued = range initk, cards = (\<lambda>g. {}), |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
37 |
roomk = initk, isin = (\<lambda>r. {}), safe = (\<lambda>r. True)\<rparr> \<in> reach" | |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
38 |
check_in: |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
39 |
"\<lbrakk>s \<in> reach; k \<notin> issued s\<rbrakk> \<Longrightarrow> |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
40 |
s\<lparr>currk := (currk s)(r := k), issued := issued s \<union> {k}, |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
41 |
cards := (cards s)(g := cards s g \<union> {(currk s r, k)}), |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
42 |
owns := (owns s)(r := Some g), safe := (safe s)(r := False)\<rparr> \<in> reach" | |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
43 |
enter_room: |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
44 |
"\<lbrakk>s \<in> reach; (k,k') \<in> cards s g; roomk s r \<in> {k,k'}\<rbrakk> \<Longrightarrow> |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
45 |
s\<lparr>isin := (isin s)(r := isin s r \<union> {g}), |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
46 |
roomk := (roomk s)(r := k'), |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
47 |
safe := (safe s)(r := owns s r = Some g \<and> isin s r = {} (* \<and> k' = currk s r *) |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
48 |
\<or> safe s r)\<rparr> \<in> reach" | |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
49 |
exit_room: |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
50 |
"\<lbrakk>s \<in> reach; g \<in> isin s r\<rbrakk> \<Longrightarrow> s\<lparr>isin := (isin s)(r := isin s r - {g})\<rparr> \<in> reach" |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
51 |
|
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
52 |
theorem safe: "s \<in> reach \<Longrightarrow> safe s r \<Longrightarrow> g \<in> isin s r \<Longrightarrow> owns s r = Some g" |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
53 |
nitpick [card room = 1, card guest = 2, card "guest option" = 3, |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
54 |
card key = 4, card state = 6, expect = genuine] |
38185 | 55 |
(* nitpick [card room = 1, card guest = 2, expect = genuine] *) |
35076
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
56 |
oops |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
57 |
|
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
diff
changeset
|
58 |
end |