35762
|
1 |
(* Title: Sequents/Washing.thy
|
|
2 |
Author: Sara Kalvala
|
|
3 |
*)
|
21426
|
4 |
|
|
5 |
theory Washing
|
|
6 |
imports ILL
|
|
7 |
begin
|
|
8 |
|
51309
|
9 |
axiomatization
|
|
10 |
dollar :: o and
|
|
11 |
quarter :: o and
|
|
12 |
loaded :: o and
|
|
13 |
dirty :: o and
|
|
14 |
wet :: o and
|
21426
|
15 |
clean :: o
|
51309
|
16 |
where
|
21426
|
17 |
change:
|
61386
|
18 |
"dollar \<turnstile> (quarter >< quarter >< quarter >< quarter)" and
|
21426
|
19 |
|
|
20 |
load1:
|
61386
|
21 |
"quarter , quarter , quarter , quarter , quarter \<turnstile> loaded" and
|
21426
|
22 |
|
|
23 |
load2:
|
61386
|
24 |
"dollar , quarter \<turnstile> loaded" and
|
21426
|
25 |
|
|
26 |
wash:
|
61386
|
27 |
"loaded , dirty \<turnstile> wet" and
|
21426
|
28 |
|
|
29 |
dry:
|
61386
|
30 |
"wet, quarter , quarter , quarter \<turnstile> clean"
|
21426
|
31 |
|
|
32 |
|
|
33 |
(* "activate" definitions for use in proof *)
|
|
34 |
|
60770
|
35 |
ML \<open>ML_Thms.bind_thms ("changeI", [@{thm context1}] RL ([@{thm change}] RLN (2,[@{thm cut}])))\<close>
|
|
36 |
ML \<open>ML_Thms.bind_thms ("load1I", [@{thm context1}] RL ([@{thm load1}] RLN (2,[@{thm cut}])))\<close>
|
|
37 |
ML \<open>ML_Thms.bind_thms ("washI", [@{thm context1}] RL ([@{thm wash}] RLN (2,[@{thm cut}])))\<close>
|
|
38 |
ML \<open>ML_Thms.bind_thms ("dryI", [@{thm context1}] RL ([@{thm dry}] RLN (2,[@{thm cut}])))\<close>
|
21426
|
39 |
|
|
40 |
(* a load of dirty clothes and two dollars gives you clean clothes *)
|
|
41 |
|
61386
|
42 |
lemma "dollar , dollar , dirty \<turnstile> clean"
|
55232
|
43 |
by (best add!: changeI load1I washI dryI)
|
21426
|
44 |
|
|
45 |
(* order of premises doesn't matter *)
|
|
46 |
|
61386
|
47 |
lemma "dollar , dirty , dollar \<turnstile> clean"
|
55232
|
48 |
by (best add!: changeI load1I washI dryI)
|
21426
|
49 |
|
|
50 |
(* alternative formulation *)
|
|
51 |
|
61386
|
52 |
lemma "dollar , dollar \<turnstile> dirty -o clean"
|
55232
|
53 |
by (best add!: changeI load1I washI dryI)
|
21426
|
54 |
|
|
55 |
end
|