| author | haftmann | 
| Sat, 19 Oct 2019 09:15:41 +0000 | |
| changeset 70903 | c550368a4e29 | 
| parent 67415 | 53d0fb1359d4 | 
| child 74641 | 6f801e1073fa | 
| 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  | 
| 45035 | 3  | 
Copyright 2010-2011  | 
| 
35076
 
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  | 
|
| 63167 | 8  | 
section \<open>Nitpick Example Based on Tobias Nipkow's Hotel Key Card  | 
9  | 
Formalization\<close>  | 
|
| 
35076
 
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  | 
|
| 
54680
 
93f6d33a754e
reverted 86e0b402994c, which was accidentally qfinish'ed and pushed
 
blanchet 
parents: 
54633 
diff
changeset
 | 
15  | 
nitpick_params [verbose, max_potential = 0, sat_solver = MiniSat_JNI,  | 
| 
42208
 
02513eb26eb7
raised timeouts further, for SML/NJ -- because of variations in machines/compilers, fixed timeouts can merely prevent non-termination, not enforce particular performance characteristics.
 
krauss 
parents: 
41360 
diff
changeset
 | 
16  | 
max_threads = 1, timeout = 240]  | 
| 
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  | 
|
| 42463 | 22  | 
type_synonym keycard = "key \<times> key"  | 
| 
35076
 
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'),  | 
| 67415 | 47  | 
   safe := (safe s)(r := owns s r = Some g \<and> isin s r = {} \<^cancel>\<open>\<and> k' = currk s r\<close>
 | 
| 
35076
 
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,  | 
| 41360 | 54  | 
card key = 4, card state = 6, show_consts, format = 2,  | 
55  | 
expect = genuine]  | 
|
| 
46082
 
1c436a920730
reintroduced failing examples now that they work again, after reintroduction of "set"
 
blanchet 
parents: 
45035 
diff
changeset
 | 
56  | 
(* nitpick [card room = 1, card guest = 2, expect = genuine] *) (* slow *)  | 
| 
35076
 
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
 
blanchet 
parents:  
diff
changeset
 | 
57  | 
oops  | 
| 
 
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
 
blanchet 
parents:  
diff
changeset
 | 
58  | 
|
| 
 
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
 
blanchet 
parents:  
diff
changeset
 | 
59  | 
end  |