author | paulson |
Fri, 03 Mar 2000 18:26:19 +0100 | |
changeset 8334 | 7896bcbd8641 |
parent 7403 | c318acb88251 |
child 9190 | b86ff604729f |
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 |
|
6570 | 32 |
Goal "Rprg : Always reach_invariant"; |
7403 | 33 |
by (always_tac 1); |
4776 | 34 |
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
|
35 |
qed "reach_invariant"; |
4776 | 36 |
|
37 |
||
38 |
(*** Fixedpoint ***) |
|
39 |
||
40 |
(*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
|
41 |
Goalw [fixedpoint_def] |
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
42 |
"fixedpoint Int reach_invariant = { %v. (init, v) : edges^* }"; |
4776 | 43 |
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
|
44 |
by (auto_tac (claset() addSIs [ext], simpset())); |
4776 | 45 |
by (blast_tac (claset() addIs [r_into_rtrancl,rtrancl_trans]) 2); |
46 |
by (etac rtrancl_induct 1); |
|
47 |
by Auto_tac; |
|
48 |
qed "fixedpoint_invariant_correct"; |
|
49 |
||
5648 | 50 |
Goalw [FP_def, fixedpoint_def, stable_def, constrains_def, Rprg_def] |
51 |
"FP Rprg <= fixedpoint"; |
|
4776 | 52 |
by Auto_tac; |
5758
27a2b36efd95
corrected auto_tac (applications of unsafe wrappers)
oheimb
parents:
5648
diff
changeset
|
53 |
by (dtac bspec 1 THEN atac 1); |
4776 | 54 |
by (asm_full_simp_tac (simpset() addsimps [Image_singleton, image_iff]) 1); |
55 |
by (dtac fun_cong 1); |
|
56 |
by Auto_tac; |
|
57 |
val lemma1 = result(); |
|
58 |
||
5648 | 59 |
Goalw [FP_def, fixedpoint_def, stable_def, constrains_def, Rprg_def] |
60 |
"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
|
61 |
by (auto_tac (claset() addSIs [ext], simpset())); |
4776 | 62 |
val lemma2 = result(); |
63 |
||
5648 | 64 |
Goal "FP Rprg = fixedpoint"; |
4776 | 65 |
by (rtac ([lemma1,lemma2] MRS equalityI) 1); |
66 |
qed "FP_fixedpoint"; |
|
67 |
||
68 |
||
69 |
(*If we haven't reached a fixedpoint then there is some edge for which u but |
|
70 |
not v holds. Progress will be proved via an ENSURES assertion that the |
|
71 |
metric will decrease for each suitable edge. A union over all edges proves |
|
72 |
a LEADSTO assertion that the metric decreases if we are not at a fixedpoint. |
|
73 |
*) |
|
74 |
||
5490 | 75 |
Goal "- fixedpoint = (UN (u,v): edges. {s. s u & ~ s v})"; |
4776 | 76 |
by (simp_tac (simpset() addsimps |
8334
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
7403
diff
changeset
|
77 |
[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
|
78 |
by Auto_tac; |
8334
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
7403
diff
changeset
|
79 |
by (rtac fun_upd_idem 1); |
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
7403
diff
changeset
|
80 |
by Auto_tac; |
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
7403
diff
changeset
|
81 |
by (force_tac (claset() addSIs [rev_bexI], |
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
7403
diff
changeset
|
82 |
simpset() addsimps [fun_upd_idem_iff]) 1); |
4776 | 83 |
qed "Compl_fixedpoint"; |
84 |
||
5069 | 85 |
Goal "A - fixedpoint = (UN (u,v): edges. A Int {s. s u & ~ s v})"; |
4776 | 86 |
by (simp_tac (simpset() addsimps [Diff_eq, Compl_fixedpoint]) 1); |
87 |
by (Blast_tac 1); |
|
88 |
qed "Diff_fixedpoint"; |
|
89 |
||
90 |
||
91 |
(*** Progress ***) |
|
92 |
||
5111 | 93 |
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
|
94 |
by (subgoal_tac "{v. ~ (s(x:=True)) v} = {v. ~ s v} - {x}" 1); |
4776 | 95 |
by Auto_tac; |
5629 | 96 |
by (asm_simp_tac (simpset() addsimps [card_Suc_Diff1]) 1); |
4776 | 97 |
qed "Suc_metric"; |
98 |
||
5111 | 99 |
Goal "~ s x ==> metric (s(x:=True)) < metric s"; |
4776 | 100 |
by (etac (Suc_metric RS subst) 1); |
101 |
by (Blast_tac 1); |
|
102 |
qed "metric_less"; |
|
103 |
AddSIs [metric_less]; |
|
104 |
||
5071
548f398d770b
Consequences of the change from [ := ] to ( := ) in theory Update.
nipkow
parents:
5069
diff
changeset
|
105 |
Goal "metric (s(y:=s x | s y)) <= metric s"; |
4776 | 106 |
by (case_tac "s x --> s y" 1); |
107 |
by (auto_tac (claset() addIs [less_imp_le], |
|
5196 | 108 |
simpset() addsimps [fun_upd_idem])); |
4776 | 109 |
qed "metric_le"; |
110 |
||
6536 | 111 |
Goal "Rprg : ((metric-``{m}) - fixedpoint) LeadsTo (metric-``(lessThan m))"; |
4776 | 112 |
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
|
113 |
by (rtac LeadsTo_UN 1); |
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5253
diff
changeset
|
114 |
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
|
115 |
by (ensures_tac "asgt a b" 1); |
5424
771a68a468cc
modified proofs for new constrains_tac and ensures_tac
paulson
parents:
5313
diff
changeset
|
116 |
by (Blast_tac 2); |
771a68a468cc
modified proofs for new constrains_tac and ensures_tac
paulson
parents:
5313
diff
changeset
|
117 |
by (full_simp_tac (simpset() addsimps [not_less_iff_le]) 1); |
6676 | 118 |
by (dtac (metric_le RS order_antisym) 1); |
5424
771a68a468cc
modified proofs for new constrains_tac and ensures_tac
paulson
parents:
5313
diff
changeset
|
119 |
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
|
120 |
simpset())); |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5253
diff
changeset
|
121 |
qed "LeadsTo_Diff_fixedpoint"; |
4776 | 122 |
|
6536 | 123 |
Goal "Rprg : (metric-``{m}) LeadsTo (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
|
124 |
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
|
125 |
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
|
126 |
by Auto_tac; |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5253
diff
changeset
|
127 |
qed "LeadsTo_Un_fixedpoint"; |
4776 | 128 |
|
129 |
||
130 |
(*Execution in any state leads to a fixedpoint (i.e. can terminate)*) |
|
6536 | 131 |
Goal "Rprg : UNIV LeadsTo fixedpoint"; |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5253
diff
changeset
|
132 |
by (rtac LessThan_induct 1); |
4776 | 133 |
by Auto_tac; |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5253
diff
changeset
|
134 |
by (rtac LeadsTo_Un_fixedpoint 1); |
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5253
diff
changeset
|
135 |
qed "LeadsTo_fixedpoint"; |
4776 | 136 |
|
6536 | 137 |
Goal "Rprg : UNIV LeadsTo { %v. (init, v) : edges^* }"; |
4776 | 138 |
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
|
139 |
by (rtac ([reach_invariant, LeadsTo_fixedpoint] |
6570 | 140 |
MRS Always_LeadsTo_weaken) 1); |
5240
bbcd79ef7cf2
Constant "invariant" and new constrains_tac, ensures_tac
paulson
parents:
5232
diff
changeset
|
141 |
by Auto_tac; |
4776 | 142 |
qed "LeadsTo_correct"; |