author | obua |
Sun, 09 May 2004 23:04:36 +0200 | |
changeset 14722 | 8e739a6eaf11 |
parent 13806 | fd40c9d9076b |
child 16417 | 9bc16273c2d4 |
permissions | -rw-r--r-- |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
1 |
(* Title: HOL/UNITY/Reachability |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
2 |
ID: $Id$ |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
3 |
Author: Tanja Vos, Cambridge University Computer Laboratory |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
4 |
Copyright 2000 University of Cambridge |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
5 |
|
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
6 |
Reachability in Graphs |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
7 |
|
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
8 |
From Chandy and Misra, "Parallel Program Design" (1989), sections 6.2 and 11.3 |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
9 |
*) |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
10 |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
11 |
theory Reachability = Detects + Reach: |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
12 |
|
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
13 |
types edge = "(vertex*vertex)" |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
14 |
|
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
15 |
record state = |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
16 |
reach :: "vertex => bool" |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
17 |
nmsg :: "edge => nat" |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
18 |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
19 |
consts REACHABLE :: "edge set" |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
20 |
root :: "vertex" |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
21 |
E :: "edge set" |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
22 |
V :: "vertex set" |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
23 |
|
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
24 |
inductive "REACHABLE" |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
25 |
intros |
13806 | 26 |
base: "v \<in> V ==> ((v,v) \<in> REACHABLE)" |
27 |
step: "((u,v) \<in> REACHABLE) & (v,w) \<in> E ==> ((u,w) \<in> REACHABLE)" |
|
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
28 |
|
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
29 |
constdefs |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
30 |
reachable :: "vertex => state set" |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
31 |
"reachable p == {s. reach s p}" |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
32 |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
33 |
nmsg_eq :: "nat => edge => state set" |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
34 |
"nmsg_eq k == %e. {s. nmsg s e = k}" |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
35 |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
36 |
nmsg_gt :: "nat => edge => state set" |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
37 |
"nmsg_gt k == %e. {s. k < nmsg s e}" |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
38 |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
39 |
nmsg_gte :: "nat => edge => state set" |
13806 | 40 |
"nmsg_gte k == %e. {s. k \<le> nmsg s e}" |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
41 |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
42 |
nmsg_lte :: "nat => edge => state set" |
13806 | 43 |
"nmsg_lte k == %e. {s. nmsg s e \<le> k}" |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
44 |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
45 |
final :: "state set" |
13806 | 46 |
"final == (\<Inter>v\<in>V. reachable v <==> {s. (root, v) \<in> REACHABLE}) \<inter> |
47 |
(INTER E (nmsg_eq 0))" |
|
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
48 |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
49 |
axioms |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
50 |
|
13806 | 51 |
Graph1: "root \<in> V" |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
52 |
|
13806 | 53 |
Graph2: "(v,w) \<in> E ==> (v \<in> V) & (w \<in> V)" |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
54 |
|
13806 | 55 |
MA1: "F \<in> Always (reachable root)" |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
56 |
|
13806 | 57 |
MA2: "v \<in> V ==> F \<in> Always (- reachable v \<union> {s. ((root,v) \<in> REACHABLE)})" |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
58 |
|
13806 | 59 |
MA3: "[|v \<in> V;w \<in> V|] ==> F \<in> Always (-(nmsg_gt 0 (v,w)) \<union> (reachable v))" |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
60 |
|
13806 | 61 |
MA4: "(v,w) \<in> E ==> |
62 |
F \<in> Always (-(reachable v) \<union> (nmsg_gt 0 (v,w)) \<union> (reachable w))" |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
63 |
|
13806 | 64 |
MA5: "[|v \<in> V; w \<in> V|] |
65 |
==> F \<in> Always (nmsg_gte 0 (v,w) \<inter> nmsg_lte (Suc 0) (v,w))" |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
66 |
|
13806 | 67 |
MA6: "[|v \<in> V|] ==> F \<in> Stable (reachable v)" |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
68 |
|
13806 | 69 |
MA6b: "[|v \<in> V;w \<in> W|] ==> F \<in> Stable (reachable v \<inter> nmsg_lte k (v,w))" |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
70 |
|
13806 | 71 |
MA7: "[|v \<in> V;w \<in> V|] ==> F \<in> UNIV LeadsTo nmsg_eq 0 (v,w)" |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
72 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
73 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
74 |
lemmas E_imp_in_V_L = Graph2 [THEN conjunct1, standard] |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
75 |
lemmas E_imp_in_V_R = Graph2 [THEN conjunct2, standard] |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
76 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
77 |
lemma lemma2: |
13806 | 78 |
"(v,w) \<in> E ==> F \<in> reachable v LeadsTo nmsg_eq 0 (v,w) \<inter> reachable v" |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
79 |
apply (rule MA7 [THEN PSP_Stable, THEN LeadsTo_weaken_L]) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
80 |
apply (rule_tac [3] MA6) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
81 |
apply (auto simp add: E_imp_in_V_L E_imp_in_V_R) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
82 |
done |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
83 |
|
13806 | 84 |
lemma Induction_base: "(v,w) \<in> E ==> F \<in> reachable v LeadsTo reachable w" |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
85 |
apply (rule MA4 [THEN Always_LeadsTo_weaken]) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
86 |
apply (rule_tac [2] lemma2) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
87 |
apply (auto simp add: nmsg_eq_def nmsg_gt_def) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
88 |
done |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
89 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
90 |
lemma REACHABLE_LeadsTo_reachable: |
13806 | 91 |
"(v,w) \<in> REACHABLE ==> F \<in> reachable v LeadsTo reachable w" |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
92 |
apply (erule REACHABLE.induct) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
93 |
apply (rule subset_imp_LeadsTo, blast) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
94 |
apply (blast intro: LeadsTo_Trans Induction_base) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
95 |
done |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
96 |
|
13806 | 97 |
lemma Detects_part1: "F \<in> {s. (root,v) \<in> REACHABLE} LeadsTo reachable v" |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
98 |
apply (rule single_LeadsTo_I) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
99 |
apply (simp split add: split_if_asm) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
100 |
apply (rule MA1 [THEN Always_LeadsToI]) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
101 |
apply (erule REACHABLE_LeadsTo_reachable [THEN LeadsTo_weaken_L], auto) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
102 |
done |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
103 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
104 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
105 |
lemma Reachability_Detected: |
13806 | 106 |
"v \<in> V ==> F \<in> (reachable v) Detects {s. (root,v) \<in> REACHABLE}" |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
107 |
apply (unfold Detects_def, auto) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
108 |
prefer 2 apply (blast intro: MA2 [THEN Always_weaken]) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
109 |
apply (rule Detects_part1 [THEN LeadsTo_weaken_L], blast) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
110 |
done |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
111 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
112 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
113 |
lemma LeadsTo_Reachability: |
13806 | 114 |
"v \<in> V ==> F \<in> UNIV LeadsTo (reachable v <==> {s. (root,v) \<in> REACHABLE})" |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
115 |
by (erule Reachability_Detected [THEN Detects_Imp_LeadstoEQ]) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
116 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
117 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
118 |
(* ------------------------------------ *) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
119 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
120 |
(* Some lemmas about <==> *) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
121 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
122 |
lemma Eq_lemma1: |
13806 | 123 |
"(reachable v <==> {s. (root,v) \<in> REACHABLE}) = |
124 |
{s. ((s \<in> reachable v) = ((root,v) \<in> REACHABLE))}" |
|
125 |
by (unfold Equality_def, blast) |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
126 |
|
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
127 |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
128 |
lemma Eq_lemma2: |
13806 | 129 |
"(reachable v <==> (if (root,v) \<in> REACHABLE then UNIV else {})) = |
130 |
{s. ((s \<in> reachable v) = ((root,v) \<in> REACHABLE))}" |
|
131 |
by (unfold Equality_def, auto) |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
132 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
133 |
(* ------------------------------------ *) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
134 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
135 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
136 |
(* Some lemmas about final (I don't need all of them!) *) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
137 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
138 |
lemma final_lemma1: |
13806 | 139 |
"(\<Inter>v \<in> V. \<Inter>w \<in> V. {s. ((s \<in> reachable v) = ((root,v) \<in> REACHABLE)) & |
140 |
s \<in> nmsg_eq 0 (v,w)}) |
|
141 |
\<subseteq> final" |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
142 |
apply (unfold final_def Equality_def, auto) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
143 |
apply (frule E_imp_in_V_R) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
144 |
apply (frule E_imp_in_V_L, blast) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
145 |
done |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
146 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
147 |
lemma final_lemma2: |
13806 | 148 |
"E\<noteq>{} |
149 |
==> (\<Inter>v \<in> V. \<Inter>e \<in> E. {s. ((s \<in> reachable v) = ((root,v) \<in> REACHABLE))} |
|
150 |
\<inter> nmsg_eq 0 e) \<subseteq> final" |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
151 |
apply (unfold final_def Equality_def) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
152 |
apply (auto split add: split_if_asm) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
153 |
apply (frule E_imp_in_V_L, blast) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
154 |
done |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
155 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
156 |
lemma final_lemma3: |
13806 | 157 |
"E\<noteq>{} |
158 |
==> (\<Inter>v \<in> V. \<Inter>e \<in> E. |
|
159 |
(reachable v <==> {s. (root,v) \<in> REACHABLE}) \<inter> nmsg_eq 0 e) |
|
160 |
\<subseteq> final" |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
161 |
apply (frule final_lemma2) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
162 |
apply (simp (no_asm_use) add: Eq_lemma2) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
163 |
done |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
164 |
|
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
165 |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
166 |
lemma final_lemma4: |
13806 | 167 |
"E\<noteq>{} |
168 |
==> (\<Inter>v \<in> V. \<Inter>e \<in> E. |
|
169 |
{s. ((s \<in> reachable v) = ((root,v) \<in> REACHABLE))} \<inter> nmsg_eq 0 e) |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
170 |
= final" |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
171 |
apply (rule subset_antisym) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
172 |
apply (erule final_lemma2) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
173 |
apply (unfold final_def Equality_def, blast) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
174 |
done |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
175 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
176 |
lemma final_lemma5: |
13806 | 177 |
"E\<noteq>{} |
178 |
==> (\<Inter>v \<in> V. \<Inter>e \<in> E. |
|
179 |
((reachable v) <==> {s. (root,v) \<in> REACHABLE}) \<inter> nmsg_eq 0 e) |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
180 |
= final" |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
181 |
apply (frule final_lemma4) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
182 |
apply (simp (no_asm_use) add: Eq_lemma2) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
183 |
done |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
184 |
|
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
185 |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
186 |
lemma final_lemma6: |
13806 | 187 |
"(\<Inter>v \<in> V. \<Inter>w \<in> V. |
188 |
(reachable v <==> {s. (root,v) \<in> REACHABLE}) \<inter> nmsg_eq 0 (v,w)) |
|
189 |
\<subseteq> final" |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
190 |
apply (simp (no_asm) add: Eq_lemma2 Int_def) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
191 |
apply (rule final_lemma1) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
192 |
done |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
193 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
194 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
195 |
lemma final_lemma7: |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
196 |
"final = |
13806 | 197 |
(\<Inter>v \<in> V. \<Inter>w \<in> V. |
198 |
((reachable v) <==> {s. (root,v) \<in> REACHABLE}) \<inter> |
|
199 |
(-{s. (v,w) \<in> E} \<union> (nmsg_eq 0 (v,w))))" |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
200 |
apply (unfold final_def) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
201 |
apply (rule subset_antisym, blast) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
202 |
apply (auto split add: split_if_asm) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
203 |
apply (blast dest: E_imp_in_V_L E_imp_in_V_R)+ |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
204 |
done |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
205 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
206 |
(* ------------------------------------ *) |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
207 |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
208 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
209 |
(* ------------------------------------ *) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
210 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
211 |
(* Stability theorems *) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
212 |
lemma not_REACHABLE_imp_Stable_not_reachable: |
13806 | 213 |
"[| v \<in> V; (root,v) \<notin> REACHABLE |] ==> F \<in> Stable (- reachable v)" |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
214 |
apply (drule MA2 [THEN AlwaysD], auto) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
215 |
done |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
216 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
217 |
lemma Stable_reachable_EQ_R: |
13806 | 218 |
"v \<in> V ==> F \<in> Stable (reachable v <==> {s. (root,v) \<in> REACHABLE})" |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
219 |
apply (simp (no_asm) add: Equality_def Eq_lemma2) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
220 |
apply (blast intro: MA6 not_REACHABLE_imp_Stable_not_reachable) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
221 |
done |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
222 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
223 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
224 |
lemma lemma4: |
13806 | 225 |
"((nmsg_gte 0 (v,w) \<inter> nmsg_lte (Suc 0) (v,w)) \<inter> |
226 |
(- nmsg_gt 0 (v,w) \<union> A)) |
|
227 |
\<subseteq> A \<union> nmsg_eq 0 (v,w)" |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
228 |
apply (unfold nmsg_gte_def nmsg_lte_def nmsg_gt_def nmsg_eq_def, auto) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
229 |
done |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
230 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
231 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
232 |
lemma lemma5: |
13806 | 233 |
"reachable v \<inter> nmsg_eq 0 (v,w) = |
234 |
((nmsg_gte 0 (v,w) \<inter> nmsg_lte (Suc 0) (v,w)) \<inter> |
|
235 |
(reachable v \<inter> nmsg_lte 0 (v,w)))" |
|
236 |
by (unfold nmsg_gte_def nmsg_lte_def nmsg_gt_def nmsg_eq_def, auto) |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
237 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
238 |
lemma lemma6: |
13806 | 239 |
"- nmsg_gt 0 (v,w) \<union> reachable v \<subseteq> nmsg_eq 0 (v,w) \<union> reachable v" |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
240 |
apply (unfold nmsg_gte_def nmsg_lte_def nmsg_gt_def nmsg_eq_def, auto) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
241 |
done |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
242 |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
243 |
lemma Always_reachable_OR_nmsg_0: |
13806 | 244 |
"[|v \<in> V; w \<in> V|] ==> F \<in> Always (reachable v \<union> nmsg_eq 0 (v,w))" |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
245 |
apply (rule Always_Int_I [OF MA5 MA3, THEN Always_weaken]) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
246 |
apply (rule_tac [5] lemma4, auto) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
247 |
done |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
248 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
249 |
lemma Stable_reachable_AND_nmsg_0: |
13806 | 250 |
"[|v \<in> V; w \<in> V|] ==> F \<in> Stable (reachable v \<inter> nmsg_eq 0 (v,w))" |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
251 |
apply (subst lemma5) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
252 |
apply (blast intro: MA5 Always_imp_Stable [THEN Stable_Int] MA6b) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
253 |
done |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
254 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
255 |
lemma Stable_nmsg_0_OR_reachable: |
13806 | 256 |
"[|v \<in> V; w \<in> V|] ==> F \<in> Stable (nmsg_eq 0 (v,w) \<union> reachable v)" |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
257 |
by (blast intro!: Always_weaken [THEN Always_imp_Stable] lemma6 MA3) |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
258 |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
259 |
lemma not_REACHABLE_imp_Stable_not_reachable_AND_nmsg_0: |
13806 | 260 |
"[| v \<in> V; w \<in> V; (root,v) \<notin> REACHABLE |] |
261 |
==> F \<in> Stable (- reachable v \<inter> nmsg_eq 0 (v,w))" |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
262 |
apply (rule Stable_Int [OF MA2 [THEN Always_imp_Stable] |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
263 |
Stable_nmsg_0_OR_reachable, |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
264 |
THEN Stable_eq]) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
265 |
prefer 4 apply blast |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
266 |
apply auto |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
267 |
done |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
268 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
269 |
lemma Stable_reachable_EQ_R_AND_nmsg_0: |
13806 | 270 |
"[| v \<in> V; w \<in> V |] |
271 |
==> F \<in> Stable ((reachable v <==> {s. (root,v) \<in> REACHABLE}) \<inter> |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
272 |
nmsg_eq 0 (v,w))" |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
273 |
by (simp add: Equality_def Eq_lemma2 Stable_reachable_AND_nmsg_0 |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
274 |
not_REACHABLE_imp_Stable_not_reachable_AND_nmsg_0) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
275 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
276 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
277 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
278 |
(* ------------------------------------ *) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
279 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
280 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
281 |
(* LeadsTo final predicate (Exercise 11.2 page 274) *) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
282 |
|
13806 | 283 |
lemma UNIV_lemma: "UNIV \<subseteq> (\<Inter>v \<in> V. UNIV)" |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
284 |
by blast |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
285 |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
286 |
lemmas UNIV_LeadsTo_completion = |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
287 |
LeadsTo_weaken_L [OF Finite_stable_completion UNIV_lemma] |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
288 |
|
13806 | 289 |
lemma LeadsTo_final_E_empty: "E={} ==> F \<in> UNIV LeadsTo final" |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
290 |
apply (unfold final_def, simp) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
291 |
apply (rule UNIV_LeadsTo_completion) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
292 |
apply safe |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
293 |
apply (erule LeadsTo_Reachability [simplified]) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
294 |
apply (drule Stable_reachable_EQ_R, simp) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
295 |
done |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
296 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
297 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
298 |
lemma Leadsto_reachability_AND_nmsg_0: |
13806 | 299 |
"[| v \<in> V; w \<in> V |] |
300 |
==> F \<in> UNIV LeadsTo |
|
301 |
((reachable v <==> {s. (root,v): REACHABLE}) \<inter> nmsg_eq 0 (v,w))" |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
302 |
apply (rule LeadsTo_Reachability [THEN LeadsTo_Trans], blast) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
303 |
apply (subgoal_tac |
13806 | 304 |
"F \<in> (reachable v <==> {s. (root,v) \<in> REACHABLE}) \<inter> |
305 |
UNIV LeadsTo (reachable v <==> {s. (root,v) \<in> REACHABLE}) \<inter> |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
306 |
nmsg_eq 0 (v,w) ") |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
307 |
apply simp |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
308 |
apply (rule PSP_Stable2) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
309 |
apply (rule MA7) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
310 |
apply (rule_tac [3] Stable_reachable_EQ_R, auto) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
311 |
done |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
312 |
|
13806 | 313 |
lemma LeadsTo_final_E_NOT_empty: "E\<noteq>{} ==> F \<in> UNIV LeadsTo final" |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
314 |
apply (rule LeadsTo_weaken_L [OF LeadsTo_weaken_R UNIV_lemma]) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
315 |
apply (rule_tac [2] final_lemma6) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
316 |
apply (rule Finite_stable_completion) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
317 |
apply blast |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
318 |
apply (rule UNIV_LeadsTo_completion) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
319 |
apply (blast intro: Stable_INT Stable_reachable_EQ_R_AND_nmsg_0 |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
320 |
Leadsto_reachability_AND_nmsg_0)+ |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
321 |
done |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
322 |
|
13806 | 323 |
lemma LeadsTo_final: "F \<in> UNIV LeadsTo final" |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
324 |
apply (case_tac "E={}") |
13806 | 325 |
apply (rule_tac [2] LeadsTo_final_E_NOT_empty) |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
326 |
apply (rule LeadsTo_final_E_empty, auto) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
327 |
done |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
328 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
329 |
(* ------------------------------------ *) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
330 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
331 |
(* Stability of final (Exercise 11.2 page 274) *) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
332 |
|
13806 | 333 |
lemma Stable_final_E_empty: "E={} ==> F \<in> Stable final" |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
334 |
apply (unfold final_def, simp) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
335 |
apply (rule Stable_INT) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
336 |
apply (drule Stable_reachable_EQ_R, simp) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
337 |
done |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
338 |
|
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
339 |
|
13806 | 340 |
lemma Stable_final_E_NOT_empty: "E\<noteq>{} ==> F \<in> Stable final" |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
341 |
apply (subst final_lemma7) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
342 |
apply (rule Stable_INT) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
343 |
apply (rule Stable_INT) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
344 |
apply (simp (no_asm) add: Eq_lemma2) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
345 |
apply safe |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
346 |
apply (rule Stable_eq) |
13806 | 347 |
apply (subgoal_tac [2] |
348 |
"({s. (s \<in> reachable v) = ((root,v) \<in> REACHABLE) } \<inter> nmsg_eq 0 (v,w)) = |
|
349 |
({s. (s \<in> reachable v) = ( (root,v) \<in> REACHABLE) } \<inter> (- UNIV \<union> nmsg_eq 0 (v,w)))") |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
350 |
prefer 2 apply blast |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
351 |
prefer 2 apply blast |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
352 |
apply (rule Stable_reachable_EQ_R_AND_nmsg_0 |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
353 |
[simplified Eq_lemma2 Collect_const]) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
354 |
apply (blast, blast) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
355 |
apply (rule Stable_eq) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
356 |
apply (rule Stable_reachable_EQ_R [simplified Eq_lemma2 Collect_const]) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
357 |
apply simp |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
358 |
apply (subgoal_tac |
13806 | 359 |
"({s. (s \<in> reachable v) = ((root,v) \<in> REACHABLE) }) = |
360 |
({s. (s \<in> reachable v) = ( (root,v) \<in> REACHABLE) } Int |
|
361 |
(- {} \<union> nmsg_eq 0 (v,w)))") |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
362 |
apply blast+ |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
363 |
done |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
364 |
|
13806 | 365 |
lemma Stable_final: "F \<in> Stable final" |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
366 |
apply (case_tac "E={}") |
13806 | 367 |
prefer 2 apply (blast intro: Stable_final_E_NOT_empty) |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
368 |
apply (blast intro: Stable_final_E_empty) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11701
diff
changeset
|
369 |
done |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
370 |
|
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
371 |
end |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
372 |