| 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:
 | 
| 51309 |     18 |   "dollar |- (quarter >< quarter >< quarter >< quarter)" and
 | 
| 21426 |     19 | 
 | 
|  |     20 |   load1:
 | 
| 51309 |     21 |   "quarter , quarter , quarter , quarter , quarter |- loaded" and
 | 
| 21426 |     22 | 
 | 
|  |     23 |   load2:
 | 
| 51309 |     24 |   "dollar , quarter |- loaded" and
 | 
| 21426 |     25 | 
 | 
|  |     26 |   wash:
 | 
| 51309 |     27 |   "loaded , dirty |- wet" and
 | 
| 21426 |     28 | 
 | 
|  |     29 |   dry:
 | 
|  |     30 |   "wet, quarter , quarter , quarter |- clean"
 | 
|  |     31 | 
 | 
|  |     32 | 
 | 
|  |     33 | (* "activate" definitions for use in proof *)
 | 
|  |     34 | 
 | 
| 39159 |     35 | ML {* bind_thms ("changeI", [@{thm context1}] RL ([@{thm change}] RLN (2,[@{thm cut}]))) *}
 | 
|  |     36 | ML {* bind_thms ("load1I", [@{thm context1}] RL ([@{thm load1}] RLN (2,[@{thm cut}]))) *}
 | 
|  |     37 | ML {* bind_thms ("washI", [@{thm context1}] RL ([@{thm wash}] RLN (2,[@{thm cut}]))) *}
 | 
|  |     38 | ML {* bind_thms ("dryI", [@{thm context1}] RL ([@{thm dry}] RLN (2,[@{thm cut}]))) *}
 | 
| 21426 |     39 | 
 | 
|  |     40 | (* a load of dirty clothes and two dollars gives you clean clothes *)
 | 
|  |     41 | 
 | 
|  |     42 | lemma "dollar , dollar , dirty |- clean"
 | 
| 39159 |     43 |   apply (tactic {* best_tac (lazy_cs add_safes (@{thms changeI} @ @{thms load1I} @ @{thms washI} @ @{thms dryI})) 1 *})
 | 
| 21426 |     44 |   done
 | 
|  |     45 | 
 | 
|  |     46 | (* order of premises doesn't matter *)
 | 
|  |     47 | 
 | 
|  |     48 | lemma "dollar , dirty , dollar |- clean"
 | 
| 39159 |     49 |   apply (tactic {* best_tac (lazy_cs add_safes (@{thms changeI} @ @{thms load1I} @ @{thms washI} @ @{thms dryI})) 1 *})
 | 
| 21426 |     50 |   done
 | 
|  |     51 | 
 | 
|  |     52 | (* alternative formulation *)
 | 
|  |     53 | 
 | 
|  |     54 | lemma "dollar , dollar |- dirty -o clean"
 | 
| 39159 |     55 |   apply (tactic {* best_tac (lazy_cs add_safes (@{thms changeI} @ @{thms load1I} @ @{thms washI} @ @{thms dryI})) 1 *})
 | 
| 21426 |     56 |   done
 | 
|  |     57 | 
 | 
|  |     58 | end
 |