author | wenzelm |
Wed, 03 Feb 1999 16:42:40 +0100 | |
changeset 6188 | c40e5ac04e3e |
parent 5648 | fe887910e32e |
child 6536 | 281d44905cab |
permissions | -rw-r--r-- |
4776 | 1 |
(* Title: HOL/UNITY/Channel |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1998 University of Cambridge |
|
5 |
||
6 |
Unordered Channel |
|
7 |
||
8 |
From Misra, "A Logic for Concurrent Programming" (1994), section 13.3 |
|
9 |
*) |
|
10 |
||
11 |
(*None represents "infinity" while Some represents proper integers*) |
|
5111 | 12 |
Goalw [minSet_def] "minSet A = Some x --> x : A"; |
4776 | 13 |
by (Simp_tac 1); |
14 |
by (fast_tac (claset() addIs [LeastI]) 1); |
|
15 |
qed_spec_mp "minSet_eq_SomeD"; |
|
16 |
||
5069 | 17 |
Goalw [minSet_def] " minSet{} = None"; |
4776 | 18 |
by (Asm_simp_tac 1); |
19 |
qed_spec_mp "minSet_empty"; |
|
20 |
Addsimps [minSet_empty]; |
|
21 |
||
5111 | 22 |
Goalw [minSet_def] "x:A ==> minSet A = Some (LEAST x. x: A)"; |
4776 | 23 |
by (ALLGOALS Asm_simp_tac); |
24 |
by (Blast_tac 1); |
|
25 |
qed_spec_mp "minSet_nonempty"; |
|
26 |
||
5648 | 27 |
Goal "F : leadsTo (minSet -`` {Some x}) (minSet -`` (Some``greaterThan x))"; |
4776 | 28 |
by (rtac leadsTo_weaken 1); |
5277
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5253
diff
changeset
|
29 |
by (rtac ([UC2, UC1] MRS psp) 1); |
4776 | 30 |
by (ALLGOALS Asm_simp_tac); |
31 |
by (Blast_tac 1); |
|
32 |
by Safe_tac; |
|
33 |
by (auto_tac (claset() addDs [minSet_eq_SomeD], |
|
5625 | 34 |
simpset() addsimps [le_def, linorder_neq_iff])); |
4776 | 35 |
qed "minSet_greaterThan"; |
36 |
||
37 |
||
38 |
(*The induction*) |
|
5648 | 39 |
Goal "F : leadsTo (UNIV-{{}}) (minSet -`` (Some``atLeast y))"; |
4776 | 40 |
by (rtac leadsTo_weaken_R 1); |
41 |
by (res_inst_tac [("l", "y"), ("f", "the o minSet"), ("B", "{}")] |
|
42 |
greaterThan_bounded_induct 1); |
|
43 |
by Safe_tac; |
|
44 |
by (ALLGOALS Asm_simp_tac); |
|
45 |
by (dtac minSet_nonempty 2); |
|
46 |
by (Asm_full_simp_tac 2); |
|
47 |
by (rtac (minSet_greaterThan RS leadsTo_weaken) 1); |
|
48 |
by Safe_tac; |
|
49 |
by (ALLGOALS Asm_full_simp_tac); |
|
50 |
by (dtac minSet_nonempty 1); |
|
51 |
by (Asm_full_simp_tac 1); |
|
52 |
val lemma = result(); |
|
53 |
||
54 |
||
5648 | 55 |
Goal "!!y::nat. F : leadsTo (UNIV-{{}}) {s. y ~: s}"; |
4776 | 56 |
by (rtac (lemma RS leadsTo_weaken_R) 1); |
57 |
by (Clarify_tac 1); |
|
58 |
by (forward_tac [minSet_nonempty] 1); |
|
59 |
by (asm_full_simp_tac (simpset() addsimps [Suc_le_eq]) 1); |
|
60 |
by (blast_tac (claset() addDs [Suc_le_lessD, not_less_Least]) 1); |
|
61 |
qed "Channel_progress"; |