| author | wenzelm | 
| Tue, 28 Jul 2009 15:05:18 +0200 | |
| changeset 32252 | fd5e4a1a4ea6 | 
| parent 26826 | fd8fdf21660f | 
| child 32960 | 69916a850301 | 
| 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  | 
|
| 26826 | 11  | 
theory Reachability imports "../Detects" Reach begin  | 
| 
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  | 
|
| 23767 | 19  | 
consts root :: "vertex"  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
20  | 
E :: "edge set"  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
21  | 
V :: "vertex set"  | 
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
22  | 
|
| 23767 | 23  | 
inductive_set REACHABLE :: "edge set"  | 
24  | 
where  | 
|
25  | 
base: "v \<in> V ==> ((v,v) \<in> REACHABLE)"  | 
|
26  | 
| 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
 | 
27  | 
|
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
28  | 
constdefs  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
29  | 
reachable :: "vertex => state set"  | 
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
30  | 
  "reachable p == {s. reach s p}"
 | 
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
31  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
32  | 
nmsg_eq :: "nat => edge => state set"  | 
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
33  | 
  "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
 | 
34  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
35  | 
nmsg_gt :: "nat => edge => state set"  | 
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
36  | 
  "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
 | 
37  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
38  | 
nmsg_gte :: "nat => edge => state set"  | 
| 13806 | 39  | 
  "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
 | 
40  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
41  | 
nmsg_lte :: "nat => edge => state set"  | 
| 13806 | 42  | 
  "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
 | 
43  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
44  | 
final :: "state set"  | 
| 13806 | 45  | 
  "final == (\<Inter>v\<in>V. reachable v <==> {s. (root, v) \<in> REACHABLE}) \<inter> 
 | 
46  | 
(INTER E (nmsg_eq 0))"  | 
|
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
47  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
48  | 
axioms  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
49  | 
|
| 13806 | 50  | 
Graph1: "root \<in> V"  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
51  | 
|
| 13806 | 52  | 
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
 | 
53  | 
|
| 13806 | 54  | 
MA1: "F \<in> Always (reachable root)"  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
55  | 
|
| 13806 | 56  | 
    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
 | 
57  | 
|
| 13806 | 58  | 
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
 | 
59  | 
|
| 13806 | 60  | 
MA4: "(v,w) \<in> E ==>  | 
61  | 
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
 | 
62  | 
|
| 13806 | 63  | 
MA5: "[|v \<in> V; w \<in> V|]  | 
64  | 
==> 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
 | 
65  | 
|
| 13806 | 66  | 
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
 | 
67  | 
|
| 13806 | 68  | 
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
 | 
69  | 
|
| 13806 | 70  | 
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
 | 
71  | 
|
| 
 
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  | 
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
 | 
74  | 
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
 | 
75  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
76  | 
lemma lemma2:  | 
| 13806 | 77  | 
"(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
 | 
78  | 
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
 | 
79  | 
apply (rule_tac [3] MA6)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
80  | 
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
 | 
81  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
82  | 
|
| 13806 | 83  | 
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
 | 
84  | 
apply (rule MA4 [THEN Always_LeadsTo_weaken])  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
85  | 
apply (rule_tac [2] lemma2)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
86  | 
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
 | 
87  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
88  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
89  | 
lemma REACHABLE_LeadsTo_reachable:  | 
| 13806 | 90  | 
"(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
 | 
91  | 
apply (erule REACHABLE.induct)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
92  | 
apply (rule subset_imp_LeadsTo, blast)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
93  | 
apply (blast intro: LeadsTo_Trans Induction_base)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
94  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
95  | 
|
| 13806 | 96  | 
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
 | 
97  | 
apply (rule single_LeadsTo_I)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
98  | 
apply (simp split add: split_if_asm)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
99  | 
apply (rule MA1 [THEN Always_LeadsToI])  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
100  | 
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
 | 
101  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
102  | 
|
| 
 
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  | 
lemma Reachability_Detected:  | 
| 13806 | 105  | 
     "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
 | 
106  | 
apply (unfold Detects_def, auto)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
107  | 
prefer 2 apply (blast intro: MA2 [THEN Always_weaken])  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
108  | 
apply (rule Detects_part1 [THEN LeadsTo_weaken_L], blast)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
109  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
110  | 
|
| 
 
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  | 
lemma LeadsTo_Reachability:  | 
| 13806 | 113  | 
     "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
 | 
114  | 
by (erule Reachability_Detected [THEN Detects_Imp_LeadstoEQ])  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
115  | 
|
| 
 
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  | 
(* Some lemmas about <==> *)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
120  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
121  | 
lemma Eq_lemma1:  | 
| 13806 | 122  | 
     "(reachable v <==> {s. (root,v) \<in> REACHABLE}) =  
 | 
123  | 
      {s. ((s \<in> reachable v) = ((root,v) \<in> REACHABLE))}"
 | 
|
124  | 
by (unfold Equality_def, blast)  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
125  | 
|
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
126  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
127  | 
lemma Eq_lemma2:  | 
| 13806 | 128  | 
     "(reachable v <==> (if (root,v) \<in> REACHABLE then UNIV else {})) =  
 | 
129  | 
      {s. ((s \<in> reachable v) = ((root,v) \<in> REACHABLE))}"
 | 
|
130  | 
by (unfold Equality_def, auto)  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
131  | 
|
| 
 
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  | 
(* 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
 | 
136  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
137  | 
lemma final_lemma1:  | 
| 13806 | 138  | 
     "(\<Inter>v \<in> V. \<Inter>w \<in> V. {s. ((s \<in> reachable v) = ((root,v) \<in> REACHABLE)) &  
 | 
139  | 
s \<in> nmsg_eq 0 (v,w)})  | 
|
140  | 
\<subseteq> final"  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
141  | 
apply (unfold final_def Equality_def, auto)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
142  | 
apply (frule E_imp_in_V_R)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
143  | 
apply (frule E_imp_in_V_L, blast)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
144  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
145  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
146  | 
lemma final_lemma2:  | 
| 13806 | 147  | 
 "E\<noteq>{}  
 | 
148  | 
  ==> (\<Inter>v \<in> V. \<Inter>e \<in> E. {s. ((s \<in> reachable v) = ((root,v) \<in> REACHABLE))}  
 | 
|
149  | 
\<inter> nmsg_eq 0 e) \<subseteq> final"  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
150  | 
apply (unfold final_def Equality_def)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
151  | 
apply (auto split add: split_if_asm)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
152  | 
apply (frule E_imp_in_V_L, blast)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
153  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
154  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
155  | 
lemma final_lemma3:  | 
| 13806 | 156  | 
     "E\<noteq>{}  
 | 
157  | 
==> (\<Inter>v \<in> V. \<Inter>e \<in> E.  | 
|
158  | 
           (reachable v <==> {s. (root,v) \<in> REACHABLE}) \<inter> nmsg_eq 0 e)  
 | 
|
159  | 
\<subseteq> final"  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
160  | 
apply (frule final_lemma2)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
161  | 
apply (simp (no_asm_use) add: Eq_lemma2)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
162  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
163  | 
|
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
164  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
165  | 
lemma final_lemma4:  | 
| 13806 | 166  | 
     "E\<noteq>{}  
 | 
167  | 
==> (\<Inter>v \<in> V. \<Inter>e \<in> E.  | 
|
168  | 
           {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
 | 
169  | 
= final"  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
170  | 
apply (rule subset_antisym)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
171  | 
apply (erule final_lemma2)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
172  | 
apply (unfold final_def Equality_def, blast)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
173  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
174  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
175  | 
lemma final_lemma5:  | 
| 13806 | 176  | 
     "E\<noteq>{}  
 | 
177  | 
==> (\<Inter>v \<in> V. \<Inter>e \<in> E.  | 
|
178  | 
           ((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
 | 
179  | 
= final"  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
180  | 
apply (frule final_lemma4)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
181  | 
apply (simp (no_asm_use) add: Eq_lemma2)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
182  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
183  | 
|
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
184  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
185  | 
lemma final_lemma6:  | 
| 13806 | 186  | 
"(\<Inter>v \<in> V. \<Inter>w \<in> V.  | 
187  | 
       (reachable v <==> {s. (root,v) \<in> REACHABLE}) \<inter> nmsg_eq 0 (v,w))  
 | 
|
188  | 
\<subseteq> final"  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
189  | 
apply (simp (no_asm) add: Eq_lemma2 Int_def)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
190  | 
apply (rule final_lemma1)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
191  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
192  | 
|
| 
 
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  | 
lemma final_lemma7:  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
195  | 
"final =  | 
| 13806 | 196  | 
(\<Inter>v \<in> V. \<Inter>w \<in> V.  | 
197  | 
       ((reachable v) <==> {s. (root,v) \<in> REACHABLE}) \<inter> 
 | 
|
198  | 
       (-{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
 | 
199  | 
apply (unfold final_def)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
200  | 
apply (rule subset_antisym, blast)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
201  | 
apply (auto split add: split_if_asm)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
202  | 
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
 | 
203  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
204  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
205  | 
(* ------------------------------------ *)  | 
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
206  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
207  | 
|
| 
 
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  | 
(* Stability theorems *)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
211  | 
lemma not_REACHABLE_imp_Stable_not_reachable:  | 
| 13806 | 212  | 
"[| 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
 | 
213  | 
apply (drule MA2 [THEN AlwaysD], auto)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
214  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
215  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
216  | 
lemma Stable_reachable_EQ_R:  | 
| 13806 | 217  | 
     "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
 | 
218  | 
apply (simp (no_asm) add: Equality_def Eq_lemma2)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
219  | 
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
 | 
220  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
221  | 
|
| 
 
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  | 
lemma lemma4:  | 
| 13806 | 224  | 
"((nmsg_gte 0 (v,w) \<inter> nmsg_lte (Suc 0) (v,w)) \<inter>  | 
225  | 
(- nmsg_gt 0 (v,w) \<union> A))  | 
|
226  | 
\<subseteq> A \<union> nmsg_eq 0 (v,w)"  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
227  | 
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
 | 
228  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
229  | 
|
| 
 
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  | 
lemma lemma5:  | 
| 13806 | 232  | 
"reachable v \<inter> nmsg_eq 0 (v,w) =  | 
233  | 
((nmsg_gte 0 (v,w) \<inter> nmsg_lte (Suc 0) (v,w)) \<inter>  | 
|
234  | 
(reachable v \<inter> nmsg_lte 0 (v,w)))"  | 
|
235  | 
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
 | 
236  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
237  | 
lemma lemma6:  | 
| 13806 | 238  | 
"- 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
 | 
239  | 
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
 | 
240  | 
done  | 
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
241  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
242  | 
lemma Always_reachable_OR_nmsg_0:  | 
| 13806 | 243  | 
"[|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
 | 
244  | 
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
 | 
245  | 
apply (rule_tac [5] lemma4, auto)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
246  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
247  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
248  | 
lemma Stable_reachable_AND_nmsg_0:  | 
| 13806 | 249  | 
"[|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
 | 
250  | 
apply (subst lemma5)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
251  | 
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
 | 
252  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
253  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
254  | 
lemma Stable_nmsg_0_OR_reachable:  | 
| 13806 | 255  | 
"[|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
 | 
256  | 
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
 | 
257  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
258  | 
lemma not_REACHABLE_imp_Stable_not_reachable_AND_nmsg_0:  | 
| 13806 | 259  | 
"[| v \<in> V; w \<in> V; (root,v) \<notin> REACHABLE |]  | 
260  | 
==> 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
 | 
261  | 
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
 | 
262  | 
Stable_nmsg_0_OR_reachable,  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
263  | 
THEN Stable_eq])  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
264  | 
prefer 4 apply blast  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
265  | 
apply auto  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
266  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
267  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
268  | 
lemma Stable_reachable_EQ_R_AND_nmsg_0:  | 
| 13806 | 269  | 
"[| v \<in> V; w \<in> V |]  | 
270  | 
      ==> 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
 | 
271  | 
nmsg_eq 0 (v,w))"  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
272  | 
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
 | 
273  | 
not_REACHABLE_imp_Stable_not_reachable_AND_nmsg_0)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
274  | 
|
| 
 
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  | 
(* LeadsTo final predicate (Exercise 11.2 page 274) *)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
281  | 
|
| 13806 | 282  | 
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
 | 
283  | 
by blast  | 
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
284  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
285  | 
lemmas UNIV_LeadsTo_completion =  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
286  | 
LeadsTo_weaken_L [OF Finite_stable_completion UNIV_lemma]  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
287  | 
|
| 13806 | 288  | 
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
 | 
289  | 
apply (unfold final_def, simp)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
290  | 
apply (rule UNIV_LeadsTo_completion)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
291  | 
apply safe  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
292  | 
apply (erule LeadsTo_Reachability [simplified])  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
293  | 
apply (drule Stable_reachable_EQ_R, simp)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
294  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
295  | 
|
| 
 
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  | 
lemma Leadsto_reachability_AND_nmsg_0:  | 
| 13806 | 298  | 
"[| v \<in> V; w \<in> V |]  | 
299  | 
==> F \<in> UNIV LeadsTo  | 
|
300  | 
             ((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
 | 
301  | 
apply (rule LeadsTo_Reachability [THEN LeadsTo_Trans], blast)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
302  | 
apply (subgoal_tac  | 
| 13806 | 303  | 
	 "F \<in> (reachable v <==> {s. (root,v) \<in> REACHABLE}) \<inter> 
 | 
304  | 
              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
 | 
305  | 
nmsg_eq 0 (v,w) ")  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
306  | 
apply simp  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
307  | 
apply (rule PSP_Stable2)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
308  | 
apply (rule MA7)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
309  | 
apply (rule_tac [3] Stable_reachable_EQ_R, auto)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
310  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
311  | 
|
| 13806 | 312  | 
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
 | 
313  | 
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
 | 
314  | 
apply (rule_tac [2] final_lemma6)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
315  | 
apply (rule Finite_stable_completion)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
316  | 
apply blast  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
317  | 
apply (rule UNIV_LeadsTo_completion)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
318  | 
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
 | 
319  | 
Leadsto_reachability_AND_nmsg_0)+  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
320  | 
done  | 
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
321  | 
|
| 13806 | 322  | 
lemma LeadsTo_final: "F \<in> UNIV LeadsTo final"  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
323  | 
apply (case_tac "E={}")
 | 
| 13806 | 324  | 
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
 | 
325  | 
apply (rule LeadsTo_final_E_empty, auto)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
326  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
327  | 
|
| 
 
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  | 
(* Stability of final (Exercise 11.2 page 274) *)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
331  | 
|
| 13806 | 332  | 
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
 | 
333  | 
apply (unfold final_def, simp)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
334  | 
apply (rule Stable_INT)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
335  | 
apply (drule Stable_reachable_EQ_R, simp)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
336  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
337  | 
|
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
338  | 
|
| 13806 | 339  | 
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
 | 
340  | 
apply (subst final_lemma7)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
341  | 
apply (rule Stable_INT)  | 
| 
 
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 (simp (no_asm) add: Eq_lemma2)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
344  | 
apply safe  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
345  | 
apply (rule Stable_eq)  | 
| 13806 | 346  | 
apply (subgoal_tac [2]  | 
347  | 
     "({s. (s \<in> reachable v) = ((root,v) \<in> REACHABLE) } \<inter> nmsg_eq 0 (v,w)) = 
 | 
|
348  | 
      ({s. (s \<in> reachable v) = ( (root,v) \<in> REACHABLE) } \<inter> (- UNIV \<union> nmsg_eq 0 (v,w)))")
 | 
|
| 26826 | 349  | 
prefer 2 apply simp  | 
| 
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  | 
apply (rule Stable_reachable_EQ_R_AND_nmsg_0  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
352  | 
[simplified Eq_lemma2 Collect_const])  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
353  | 
apply (blast, blast)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
354  | 
apply (rule Stable_eq)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
355  | 
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
 | 
356  | 
apply simp  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
357  | 
apply (subgoal_tac  | 
| 13806 | 358  | 
        "({s. (s \<in> reachable v) = ((root,v) \<in> REACHABLE) }) = 
 | 
359  | 
         ({s. (s \<in> reachable v) = ( (root,v) \<in> REACHABLE) } Int
 | 
|
360  | 
              (- {} \<union> nmsg_eq 0 (v,w)))")
 | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
361  | 
apply blast+  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
362  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
363  | 
|
| 13806 | 364  | 
lemma Stable_final: "F \<in> Stable final"  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
365  | 
apply (case_tac "E={}")
 | 
| 13806 | 366  | 
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
 | 
367  | 
apply (blast intro: Stable_final_E_empty)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11701 
diff
changeset
 | 
368  | 
done  | 
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
369  | 
|
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
370  | 
end  | 
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
371  |