| author | paulson |
| Mon, 19 Oct 1998 11:26:46 +0200 | |
| changeset 5669 | f5d9caafc3bd |
| parent 5648 | fe887910e32e |
| child 5758 | 27a2b36efd95 |
| permissions | -rw-r--r-- |
| 4776 | 1 |
(* Title: HOL/UNITY/Reach.thy |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1998 University of Cambridge |
|
5 |
||
6 |
Reachability in Directed Graphs. From Chandy and Misra, section 6.4. |
|
|
4896
4727272f3db6
New syntax for function update; moved to main HOL directory
paulson
parents:
4776
diff
changeset
|
7 |
[ this example took only four days!] |
| 4776 | 8 |
*) |
9 |
||
10 |
(*TO SIMPDATA.ML?? FOR CLASET?? *) |
|
11 |
val major::prems = goal thy |
|
12 |
"[| if P then Q else R; \ |
|
13 |
\ [| P; Q |] ==> S; \ |
|
14 |
\ [| ~ P; R |] ==> S |] ==> S"; |
|
15 |
by (cut_facts_tac [major] 1); |
|
16 |
by (blast_tac (claset() addSDs [if_bool_eq_disj RS iffD1] addIs prems) 1); |
|
17 |
qed "ifE"; |
|
18 |
||
19 |
AddSEs [ifE]; |
|
20 |
||
21 |
||
| 5648 | 22 |
Addsimps [Rprg_def RS def_prg_Init]; |
23 |
program_defs_ref := [Rprg_def]; |
|
| 4776 | 24 |
|
|
5426
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
25 |
Addsimps [simp_of_act asgt_def]; |
| 4776 | 26 |
|
27 |
(*All vertex sets are finite*) |
|
28 |
AddIffs [[subset_UNIV, finite_graph] MRS finite_subset]; |
|
29 |
||
|
5426
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
30 |
Addsimps [simp_of_set reach_invariant_def]; |
| 4776 | 31 |
|
| 5648 | 32 |
Goal "Rprg : Invariant reach_invariant"; |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5253
diff
changeset
|
33 |
by (rtac InvariantI 1); |
|
5240
bbcd79ef7cf2
Constant "invariant" and new constrains_tac, ensures_tac
paulson
parents:
5232
diff
changeset
|
34 |
by Auto_tac; (*for the initial state; proof of stability remains*) |
|
5426
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
35 |
by (constrains_tac 1); |
| 4776 | 36 |
by (blast_tac (claset() addIs [r_into_rtrancl,rtrancl_trans]) 1); |
|
5240
bbcd79ef7cf2
Constant "invariant" and new constrains_tac, ensures_tac
paulson
parents:
5232
diff
changeset
|
37 |
qed "reach_invariant"; |
| 4776 | 38 |
|
39 |
||
40 |
(*** Fixedpoint ***) |
|
41 |
||
42 |
(*If it reaches a fixedpoint, it has found a solution*) |
|
|
5426
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
43 |
Goalw [fixedpoint_def] |
|
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
44 |
"fixedpoint Int reach_invariant = { %v. (init, v) : edges^* }";
|
| 4776 | 45 |
by (rtac equalityI 1); |
|
5426
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
46 |
by (auto_tac (claset() addSIs [ext], simpset())); |
| 4776 | 47 |
by (blast_tac (claset() addIs [r_into_rtrancl,rtrancl_trans]) 2); |
48 |
by (etac rtrancl_induct 1); |
|
49 |
by Auto_tac; |
|
50 |
qed "fixedpoint_invariant_correct"; |
|
51 |
||
| 5648 | 52 |
Goalw [FP_def, fixedpoint_def, stable_def, constrains_def, Rprg_def] |
53 |
"FP Rprg <= fixedpoint"; |
|
| 4776 | 54 |
by Auto_tac; |
55 |
by (asm_full_simp_tac (simpset() addsimps [Image_singleton, image_iff]) 1); |
|
56 |
by (dtac fun_cong 1); |
|
57 |
by Auto_tac; |
|
58 |
val lemma1 = result(); |
|
59 |
||
| 5648 | 60 |
Goalw [FP_def, fixedpoint_def, stable_def, constrains_def, Rprg_def] |
61 |
"fixedpoint <= FP Rprg"; |
|
|
5426
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
62 |
by (auto_tac (claset() addSIs [ext], simpset())); |
| 4776 | 63 |
val lemma2 = result(); |
64 |
||
| 5648 | 65 |
Goal "FP Rprg = fixedpoint"; |
| 4776 | 66 |
by (rtac ([lemma1,lemma2] MRS equalityI) 1); |
67 |
qed "FP_fixedpoint"; |
|
68 |
||
69 |
||
70 |
(*If we haven't reached a fixedpoint then there is some edge for which u but |
|
71 |
not v holds. Progress will be proved via an ENSURES assertion that the |
|
72 |
metric will decrease for each suitable edge. A union over all edges proves |
|
73 |
a LEADSTO assertion that the metric decreases if we are not at a fixedpoint. |
|
74 |
*) |
|
75 |
||
| 5490 | 76 |
Goal "- fixedpoint = (UN (u,v): edges. {s. s u & ~ s v})";
|
| 4776 | 77 |
by (simp_tac (simpset() addsimps |
| 5648 | 78 |
([Compl_FP, UN_UN_flatten, FP_fixedpoint RS sym, Rprg_def])) 1); |
|
5426
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
79 |
by Auto_tac; |
| 5196 | 80 |
by (rtac fun_upd_idem 1); |
| 5521 | 81 |
by (auto_tac (claset(),simpset() addsimps [fun_upd_idem_iff])); |
| 4776 | 82 |
qed "Compl_fixedpoint"; |
83 |
||
| 5069 | 84 |
Goal "A - fixedpoint = (UN (u,v): edges. A Int {s. s u & ~ s v})";
|
| 4776 | 85 |
by (simp_tac (simpset() addsimps [Diff_eq, Compl_fixedpoint]) 1); |
86 |
by (Blast_tac 1); |
|
87 |
qed "Diff_fixedpoint"; |
|
88 |
||
89 |
||
90 |
(*** Progress ***) |
|
91 |
||
| 5111 | 92 |
Goalw [metric_def] "~ s x ==> Suc (metric (s(x:=True))) = metric s"; |
|
5071
548f398d770b
Consequences of the change from [ := ] to ( := ) in theory Update.
nipkow
parents:
5069
diff
changeset
|
93 |
by (subgoal_tac "{v. ~ (s(x:=True)) v} = {v. ~ s v} - {x}" 1);
|
| 4776 | 94 |
by Auto_tac; |
| 5629 | 95 |
by (asm_simp_tac (simpset() addsimps [card_Suc_Diff1]) 1); |
| 4776 | 96 |
qed "Suc_metric"; |
97 |
||
| 5111 | 98 |
Goal "~ s x ==> metric (s(x:=True)) < metric s"; |
| 4776 | 99 |
by (etac (Suc_metric RS subst) 1); |
100 |
by (Blast_tac 1); |
|
101 |
qed "metric_less"; |
|
102 |
AddSIs [metric_less]; |
|
103 |
||
|
5071
548f398d770b
Consequences of the change from [ := ] to ( := ) in theory Update.
nipkow
parents:
5069
diff
changeset
|
104 |
Goal "metric (s(y:=s x | s y)) <= metric s"; |
| 4776 | 105 |
by (case_tac "s x --> s y" 1); |
106 |
by (auto_tac (claset() addIs [less_imp_le], |
|
| 5196 | 107 |
simpset() addsimps [fun_upd_idem])); |
| 4776 | 108 |
qed "metric_le"; |
109 |
||
| 5648 | 110 |
Goal "Rprg : LeadsTo ((metric-``{m}) - fixedpoint) (metric-``(lessThan m))";
|
| 4776 | 111 |
by (simp_tac (simpset() addsimps [Diff_fixedpoint]) 1); |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5253
diff
changeset
|
112 |
by (rtac LeadsTo_UN 1); |
|
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5253
diff
changeset
|
113 |
by Auto_tac; |
|
5426
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
114 |
by (ensures_tac "asgt a b" 1); |
|
5424
771a68a468cc
modified proofs for new constrains_tac and ensures_tac
paulson
parents:
5313
diff
changeset
|
115 |
by (Blast_tac 2); |
|
771a68a468cc
modified proofs for new constrains_tac and ensures_tac
paulson
parents:
5313
diff
changeset
|
116 |
by (full_simp_tac (simpset() addsimps [not_less_iff_le]) 1); |
| 5536 | 117 |
by (dtac (metric_le RS le_anti_sym) 1); |
|
5424
771a68a468cc
modified proofs for new constrains_tac and ensures_tac
paulson
parents:
5313
diff
changeset
|
118 |
by (auto_tac (claset() addEs [less_not_refl3 RSN (2, rev_notE)], |
|
771a68a468cc
modified proofs for new constrains_tac and ensures_tac
paulson
parents:
5313
diff
changeset
|
119 |
simpset())); |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5253
diff
changeset
|
120 |
qed "LeadsTo_Diff_fixedpoint"; |
| 4776 | 121 |
|
| 5648 | 122 |
Goal "Rprg : LeadsTo (metric-``{m}) (metric-``(lessThan m) Un fixedpoint)";
|
|
5426
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
123 |
by (rtac ([LeadsTo_Diff_fixedpoint RS LeadsTo_weaken_R, |
|
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
124 |
subset_imp_LeadsTo] MRS LeadsTo_Diff) 1); |
|
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
125 |
by Auto_tac; |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5253
diff
changeset
|
126 |
qed "LeadsTo_Un_fixedpoint"; |
| 4776 | 127 |
|
128 |
||
129 |
(*Execution in any state leads to a fixedpoint (i.e. can terminate)*) |
|
| 5648 | 130 |
Goal "Rprg : LeadsTo UNIV fixedpoint"; |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5253
diff
changeset
|
131 |
by (rtac LessThan_induct 1); |
| 4776 | 132 |
by Auto_tac; |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5253
diff
changeset
|
133 |
by (rtac LeadsTo_Un_fixedpoint 1); |
|
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5253
diff
changeset
|
134 |
qed "LeadsTo_fixedpoint"; |
| 4776 | 135 |
|
| 5648 | 136 |
Goal "Rprg : LeadsTo UNIV { %v. (init, v) : edges^* }";
|
| 4776 | 137 |
by (stac (fixedpoint_invariant_correct RS sym) 1); |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5253
diff
changeset
|
138 |
by (rtac ([reach_invariant, LeadsTo_fixedpoint] |
|
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5253
diff
changeset
|
139 |
MRS Invariant_LeadsTo_weaken) 1); |
|
5240
bbcd79ef7cf2
Constant "invariant" and new constrains_tac, ensures_tac
paulson
parents:
5232
diff
changeset
|
140 |
by Auto_tac; |
| 4776 | 141 |
qed "LeadsTo_correct"; |