src/HOL/UNITY/Simple/Reachability.thy
author wenzelm
Tue, 13 Mar 2012 22:49:02 +0100
changeset 46911 6d2a2f0e904e
parent 46008 c296c75f4cf4
child 62390 842917225d56
permissions -rw-r--r--
tuned proofs;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37936
1e4c5015a72e updated some headers;
wenzelm
parents: 35416
diff changeset
     1
(*  Title:      HOL/UNITY/Simple/Reachability.thy
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     2
    Author:     Tanja Vos, Cambridge University Computer Laboratory
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     3
    Copyright   2000  University of Cambridge
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     4
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 26826
diff changeset
     5
Reachability in Graphs.
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     6
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 26826
diff changeset
     7
From Chandy and Misra, "Parallel Program Design" (1989), sections 6.2
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 26826
diff changeset
     8
and 11.3.
11195
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
fd8fdf21660f - Tuned imports
berghofe
parents: 23767
diff changeset
    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
42463
f270e3e18be5 modernized specifications;
wenzelm
parents: 37936
diff changeset
    13
type_synonym edge = "vertex * vertex"
11195
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
7272a839ccd9 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    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
7272a839ccd9 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    23
inductive_set REACHABLE :: "edge set"
7272a839ccd9 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    24
  where
7272a839ccd9 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    25
    base: "v \<in> V ==> ((v,v) \<in> REACHABLE)"
7272a839ccd9 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    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
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
    28
definition reachable :: "vertex => state set" where
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    29
  "reachable p == {s. reach s p}"
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    30
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
    31
definition nmsg_eq :: "nat => edge  => state set" where
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    32
  "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
    33
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
    34
definition nmsg_gt :: "nat => edge  => state set" where
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    35
  "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
    36
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
    37
definition nmsg_gte :: "nat => edge => state set" where
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    38
  "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
    39
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
    40
definition nmsg_lte  :: "nat => edge => state set" where
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    41
  "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
    42
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
    43
definition final :: "state set" where
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    44
  "final == (\<Inter>v\<in>V. reachable v <==> {s. (root, v) \<in> REACHABLE}) \<inter> 
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    45
            (INTER E (nmsg_eq 0))"
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    46
45827
66c68453455c modernized specifications;
wenzelm
parents: 45605
diff changeset
    47
axiomatization
66c68453455c modernized specifications;
wenzelm
parents: 45605
diff changeset
    48
where
66c68453455c modernized specifications;
wenzelm
parents: 45605
diff changeset
    49
    Graph1: "root \<in> V" and
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    50
45827
66c68453455c modernized specifications;
wenzelm
parents: 45605
diff changeset
    51
    Graph2: "(v,w) \<in> E ==> (v \<in> V) & (w \<in> V)" and
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    52
45827
66c68453455c modernized specifications;
wenzelm
parents: 45605
diff changeset
    53
    MA1:  "F \<in> Always (reachable root)" and
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    54
45827
66c68453455c modernized specifications;
wenzelm
parents: 45605
diff changeset
    55
    MA2:  "v \<in> V ==> F \<in> Always (- reachable v \<union> {s. ((root,v) \<in> REACHABLE)})" and
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    56
45827
66c68453455c modernized specifications;
wenzelm
parents: 45605
diff changeset
    57
    MA3:  "[|v \<in> V;w \<in> V|] ==> F \<in> Always (-(nmsg_gt 0 (v,w)) \<union> (reachable v))" and
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    58
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    59
    MA4:  "(v,w) \<in> E ==> 
45827
66c68453455c modernized specifications;
wenzelm
parents: 45605
diff changeset
    60
           F \<in> Always (-(reachable v) \<union> (nmsg_gt 0 (v,w)) \<union> (reachable w))" and
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    61
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    62
    MA5:  "[|v \<in> V; w \<in> V|] 
45827
66c68453455c modernized specifications;
wenzelm
parents: 45605
diff changeset
    63
           ==> F \<in> Always (nmsg_gte 0 (v,w) \<inter> nmsg_lte (Suc 0) (v,w))" and
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    64
45827
66c68453455c modernized specifications;
wenzelm
parents: 45605
diff changeset
    65
    MA6:  "[|v \<in> V|] ==> F \<in> Stable (reachable v)" and
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    66
45827
66c68453455c modernized specifications;
wenzelm
parents: 45605
diff changeset
    67
    MA6b: "[|v \<in> V;w \<in> W|] ==> F \<in> Stable (reachable v \<inter> nmsg_lte k (v,w))" and
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    68
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    69
    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
    70
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    71
45605
a89b4bc311a5 eliminated obsolete "standard";
wenzelm
parents: 42463
diff changeset
    72
lemmas E_imp_in_V_L = Graph2 [THEN conjunct1]
a89b4bc311a5 eliminated obsolete "standard";
wenzelm
parents: 42463
diff changeset
    73
lemmas E_imp_in_V_R = Graph2 [THEN conjunct2]
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    74
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    75
lemma lemma2:
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    76
     "(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
    77
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
    78
apply (rule_tac [3] MA6)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    79
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
    80
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    81
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    82
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
    83
apply (rule MA4 [THEN Always_LeadsTo_weaken])
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    84
apply (rule_tac [2] lemma2)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    85
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
    86
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    87
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    88
lemma REACHABLE_LeadsTo_reachable:
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    89
     "(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
    90
apply (erule REACHABLE.induct)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    91
apply (rule subset_imp_LeadsTo, blast)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    92
apply (blast intro: LeadsTo_Trans Induction_base)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    93
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    94
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    95
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
    96
apply (rule single_LeadsTo_I)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    97
apply (simp split add: split_if_asm)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    98
apply (rule MA1 [THEN Always_LeadsToI])
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
    99
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
   100
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   101
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
lemma Reachability_Detected: 
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   104
     "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
   105
apply (unfold Detects_def, auto)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   106
 prefer 2 apply (blast intro: MA2 [THEN Always_weaken])
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   107
apply (rule Detects_part1 [THEN LeadsTo_weaken_L], blast)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   108
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   109
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
lemma LeadsTo_Reachability:
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   112
     "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
   113
by (erule Reachability_Detected [THEN Detects_Imp_LeadstoEQ])
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   114
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
(* Some lemmas about <==> *)
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
lemma Eq_lemma1: 
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   121
     "(reachable v <==> {s. (root,v) \<in> REACHABLE}) =  
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   122
      {s. ((s \<in> reachable v) = ((root,v) \<in> REACHABLE))}"
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   123
by (unfold Equality_def, blast)
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   124
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   125
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   126
lemma Eq_lemma2: 
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   127
     "(reachable v <==> (if (root,v) \<in> REACHABLE then UNIV else {})) =  
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   128
      {s. ((s \<in> reachable v) = ((root,v) \<in> REACHABLE))}"
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   129
by (unfold Equality_def, auto)
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   130
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
(* 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
   135
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   136
lemma final_lemma1: 
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   137
     "(\<Inter>v \<in> V. \<Inter>w \<in> V. {s. ((s \<in> reachable v) = ((root,v) \<in> REACHABLE)) &  
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   138
                              s \<in> nmsg_eq 0 (v,w)})  
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   139
      \<subseteq> final"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   140
apply (unfold final_def Equality_def, auto)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   141
apply (frule E_imp_in_V_R)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   142
apply (frule E_imp_in_V_L, blast)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   143
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   144
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   145
lemma final_lemma2: 
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   146
 "E\<noteq>{}  
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   147
  ==> (\<Inter>v \<in> V. \<Inter>e \<in> E. {s. ((s \<in> reachable v) = ((root,v) \<in> REACHABLE))}  
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   148
                           \<inter> nmsg_eq 0 e)    \<subseteq>  final"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   149
apply (unfold final_def Equality_def)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   150
apply (auto split add: split_if_asm)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   151
apply (frule E_imp_in_V_L, blast)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   152
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   153
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   154
lemma final_lemma3:
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   155
     "E\<noteq>{}  
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   156
      ==> (\<Inter>v \<in> V. \<Inter>e \<in> E.  
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   157
           (reachable v <==> {s. (root,v) \<in> REACHABLE}) \<inter> nmsg_eq 0 e)  
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   158
          \<subseteq> final"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   159
apply (frule final_lemma2)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   160
apply (simp (no_asm_use) add: Eq_lemma2)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   161
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   162
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   163
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   164
lemma final_lemma4:
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   165
     "E\<noteq>{}  
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   166
      ==> (\<Inter>v \<in> V. \<Inter>e \<in> E.  
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   167
           {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
   168
          = final"
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   169
apply (rule subset_antisym)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   170
apply (erule final_lemma2)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   171
apply (unfold final_def Equality_def, blast)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   172
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   173
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   174
lemma final_lemma5:
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   175
     "E\<noteq>{}  
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   176
      ==> (\<Inter>v \<in> V. \<Inter>e \<in> E.  
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   177
           ((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
   178
          = final"
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   179
apply (frule final_lemma4)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   180
apply (simp (no_asm_use) add: Eq_lemma2)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   181
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   182
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   183
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   184
lemma final_lemma6:
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   185
     "(\<Inter>v \<in> V. \<Inter>w \<in> V.  
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   186
       (reachable v <==> {s. (root,v) \<in> REACHABLE}) \<inter> nmsg_eq 0 (v,w))  
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   187
      \<subseteq> final"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   188
apply (simp (no_asm) add: Eq_lemma2 Int_def)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   189
apply (rule final_lemma1)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   190
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   191
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
lemma final_lemma7: 
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   194
     "final =  
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   195
      (\<Inter>v \<in> V. \<Inter>w \<in> V.  
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   196
       ((reachable v) <==> {s. (root,v) \<in> REACHABLE}) \<inter> 
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   197
       (-{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
   198
apply (unfold final_def)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   199
apply (rule subset_antisym, blast)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   200
apply (auto split add: split_if_asm)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   201
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
   202
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   203
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   204
(* ------------------------------------ *)
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   205
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   206
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
(* Stability theorems *)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   210
lemma not_REACHABLE_imp_Stable_not_reachable:
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   211
     "[| 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
   212
apply (drule MA2 [THEN AlwaysD], auto)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   213
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   214
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   215
lemma Stable_reachable_EQ_R:
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   216
     "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
   217
apply (simp (no_asm) add: Equality_def Eq_lemma2)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   218
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
   219
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   220
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
lemma lemma4: 
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   223
     "((nmsg_gte 0 (v,w) \<inter> nmsg_lte (Suc 0) (v,w)) \<inter> 
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   224
       (- nmsg_gt 0 (v,w) \<union> A))  
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   225
      \<subseteq> A \<union> nmsg_eq 0 (v,w)"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   226
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
   227
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   228
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
lemma lemma5: 
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   231
     "reachable v \<inter> nmsg_eq 0 (v,w) =  
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   232
      ((nmsg_gte 0 (v,w) \<inter> nmsg_lte (Suc 0) (v,w)) \<inter> 
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   233
       (reachable v \<inter> nmsg_lte 0 (v,w)))"
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   234
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
   235
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   236
lemma lemma6: 
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   237
     "- 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
   238
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
   239
done
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   240
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   241
lemma Always_reachable_OR_nmsg_0:
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   242
     "[|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
   243
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
   244
apply (rule_tac [5] lemma4, auto)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   245
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   246
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   247
lemma Stable_reachable_AND_nmsg_0:
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   248
     "[|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
   249
apply (subst lemma5)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   250
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
   251
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   252
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   253
lemma Stable_nmsg_0_OR_reachable:
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   254
     "[|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
   255
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
   256
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   257
lemma not_REACHABLE_imp_Stable_not_reachable_AND_nmsg_0:
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   258
     "[| v \<in> V; w \<in> V; (root,v) \<notin> REACHABLE |]  
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   259
      ==> 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
   260
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
   261
                           Stable_nmsg_0_OR_reachable, 
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   262
            THEN Stable_eq])
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   263
   prefer 4 apply blast
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   264
apply auto
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   265
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   266
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   267
lemma Stable_reachable_EQ_R_AND_nmsg_0:
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   268
     "[| v \<in> V; w \<in> V |]  
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   269
      ==> 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
   270
                      nmsg_eq 0 (v,w))"
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   271
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
   272
              not_REACHABLE_imp_Stable_not_reachable_AND_nmsg_0)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   273
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
(* LeadsTo final predicate (Exercise 11.2 page 274) *)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   280
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   281
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
   282
by blast
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   283
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   284
lemmas UNIV_LeadsTo_completion = 
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   285
    LeadsTo_weaken_L [OF Finite_stable_completion UNIV_lemma]
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   286
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   287
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
   288
apply (unfold final_def, simp)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   289
apply (rule UNIV_LeadsTo_completion)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   290
  apply safe
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   291
 apply (erule LeadsTo_Reachability [simplified])
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   292
apply (drule Stable_reachable_EQ_R, simp)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   293
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   294
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
lemma Leadsto_reachability_AND_nmsg_0:
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   297
     "[| v \<in> V; w \<in> V |]  
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   298
      ==> F \<in> UNIV LeadsTo  
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   299
             ((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
   300
apply (rule LeadsTo_Reachability [THEN LeadsTo_Trans], blast)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   301
apply (subgoal_tac
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 26826
diff changeset
   302
         "F \<in> (reachable v <==> {s. (root,v) \<in> REACHABLE}) \<inter> 
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   303
              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
   304
              nmsg_eq 0 (v,w) ")
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   305
apply simp
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   306
apply (rule PSP_Stable2)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   307
apply (rule MA7)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   308
apply (rule_tac [3] Stable_reachable_EQ_R, auto)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   309
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   310
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   311
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
   312
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
   313
apply (rule_tac [2] final_lemma6)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   314
apply (rule Finite_stable_completion)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   315
  apply blast
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   316
 apply (rule UNIV_LeadsTo_completion)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   317
   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
   318
                    Leadsto_reachability_AND_nmsg_0)+
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   319
done
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   320
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   321
lemma LeadsTo_final: "F \<in> UNIV LeadsTo final"
46911
6d2a2f0e904e tuned proofs;
wenzelm
parents: 46008
diff changeset
   322
apply (cases "E={}")
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   323
 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
   324
apply (rule LeadsTo_final_E_empty, auto)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   325
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   326
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
(* Stability of final (Exercise 11.2 page 274) *)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   330
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   331
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
   332
apply (unfold final_def, simp)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   333
apply (rule Stable_INT)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   334
apply (drule Stable_reachable_EQ_R, simp)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   335
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   336
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   337
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   338
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
   339
apply (subst final_lemma7)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   340
apply (rule Stable_INT)
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 (simp (no_asm) add: Eq_lemma2)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   343
apply safe
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   344
apply (rule Stable_eq)
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   345
apply (subgoal_tac [2]
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   346
     "({s. (s \<in> reachable v) = ((root,v) \<in> REACHABLE) } \<inter> nmsg_eq 0 (v,w)) = 
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   347
      ({s. (s \<in> reachable v) = ( (root,v) \<in> REACHABLE) } \<inter> (- UNIV \<union> nmsg_eq 0 (v,w)))")
46008
c296c75f4cf4 reverted some changes for set->predicate transition, according to "hg log -u berghofe -r Isabelle2007:Isabelle2008";
wenzelm
parents: 45827
diff changeset
   348
prefer 2 apply blast
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   349
prefer 2 apply blast 
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   350
apply (rule Stable_reachable_EQ_R_AND_nmsg_0
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   351
            [simplified Eq_lemma2 Collect_const])
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   352
apply (blast, blast)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   353
apply (rule Stable_eq)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   354
 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
   355
 apply simp
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   356
apply (subgoal_tac 
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   357
        "({s. (s \<in> reachable v) = ((root,v) \<in> REACHABLE) }) = 
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   358
         ({s. (s \<in> reachable v) = ( (root,v) \<in> REACHABLE) } Int
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   359
              (- {} \<union> nmsg_eq 0 (v,w)))")
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   360
apply blast+
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   361
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   362
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   363
lemma Stable_final: "F \<in> Stable final"
46911
6d2a2f0e904e tuned proofs;
wenzelm
parents: 46008
diff changeset
   364
apply (cases "E={}")
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   365
 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
   366
apply (blast intro: Stable_final_E_empty)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11701
diff changeset
   367
done
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   368
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   369
end
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   370