src/HOL/UNITY/Simple/Reach.thy
author haftmann
Mon, 01 Mar 2010 13:40:23 +0100
changeset 35416 d8d7d1b785af
parent 32960 69916a850301
child 36866 426d5781bb25
permissions -rw-r--r--
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/Reach.thy
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     3
    Copyright   1998  University of Cambridge
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     4
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     5
Reachability in Directed Graphs.  From Chandy and Misra, section 6.4.
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 26825
diff changeset
     6
[this example took only four days!]
11195
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
18556
dc39832e9280 added explicit paths to required theories
paulson
parents: 16796
diff changeset
     9
theory Reach imports "../UNITY_Main" begin
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    10
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    11
typedecl vertex
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    12
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    13
types    state = "vertex=>bool"
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
consts
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    16
  init ::  "vertex"
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    17
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    18
  edges :: "(vertex*vertex) set"
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    19
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
    20
definition asgt :: "[vertex,vertex] => (state*state) set" where
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    21
    "asgt u v == {(s,s'). s' = s(v:= s u | s v)}"
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    22
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
    23
definition Rprg :: "state program" where
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13806
diff changeset
    24
    "Rprg == mk_total_program ({%v. v=init}, \<Union>(u,v)\<in>edges. {asgt u v}, UNIV)"
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    25
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
    26
definition reach_invariant :: "state set" where
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    27
    "reach_invariant == {s. (\<forall>v. s v --> (init, v) \<in> edges^*) & s init}"
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    28
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
    29
definition fixedpoint :: "state set" where
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    30
    "fixedpoint == {s. \<forall>(u,v)\<in>edges. s u --> s v}"
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    31
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
    32
definition metric :: "state => nat" where
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    33
    "metric s == card {v. ~ s v}"
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: 12338
diff changeset
    35
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    36
text{**We assume that the set of vertices is finite*}
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    37
axioms 
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    38
  finite_graph:  "finite (UNIV :: vertex set)"
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    39
  
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    40
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    41
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    42
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    43
(*TO SIMPDATA.ML??  FOR CLASET??  *)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    44
lemma ifE [elim!]:
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    45
    "[| if P then Q else R;     
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    46
        [| P;   Q |] ==> S;     
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    47
        [| ~ P; R |] ==> S |] ==> S"
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    48
by (simp split: split_if_asm) 
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    49
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    50
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    51
declare Rprg_def [THEN def_prg_Init, simp]
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    52
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    53
declare asgt_def [THEN def_act_simp,simp]
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    54
14150
9a23e4eb5eb3 A document for UNITY
paulson
parents: 13812
diff changeset
    55
text{*All vertex sets are finite*}
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    56
declare finite_subset [OF subset_UNIV finite_graph, iff]
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    57
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    58
declare reach_invariant_def [THEN def_set_simp, simp]
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    59
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    60
lemma reach_invariant: "Rprg \<in> Always reach_invariant"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    61
apply (rule AlwaysI, force) 
16184
80617b8d33c5 renamed "constrains" to "safety" to avoid keyword clash
paulson
parents: 15097
diff changeset
    62
apply (unfold Rprg_def, safety) 
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    63
apply (blast intro: rtrancl_trans)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    64
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    65
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    66
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    67
(*** Fixedpoint ***)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    68
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    69
(*If it reaches a fixedpoint, it has found a solution*)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    70
lemma fixedpoint_invariant_correct: 
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    71
     "fixedpoint \<inter> reach_invariant = { %v. (init, v) \<in> edges^* }"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    72
apply (unfold fixedpoint_def)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    73
apply (rule equalityI)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    74
apply (auto intro!: ext)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    75
apply (erule rtrancl_induct, auto)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    76
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    77
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    78
lemma lemma1: 
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    79
     "FP Rprg \<subseteq> fixedpoint"
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13806
diff changeset
    80
apply (simp add: FP_def fixedpoint_def Rprg_def mk_total_program_def)
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13806
diff changeset
    81
apply (auto simp add: stable_def constrains_def)
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    82
apply (drule bspec, assumption)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    83
apply (simp add: Image_singleton image_iff)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    84
apply (drule fun_cong, auto)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    85
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    86
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    87
lemma lemma2: 
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    88
     "fixedpoint \<subseteq> FP Rprg"
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13806
diff changeset
    89
apply (simp add: FP_def fixedpoint_def Rprg_def mk_total_program_def)
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13806
diff changeset
    90
apply (auto intro!: ext simp add: stable_def constrains_def)
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    91
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    92
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    93
lemma FP_fixedpoint: "FP Rprg = fixedpoint"
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    94
by (rule equalityI [OF lemma1 lemma2])
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    95
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    96
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    97
(*If we haven't reached a fixedpoint then there is some edge for which u but
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    98
  not v holds.  Progress will be proved via an ENSURES assertion that the
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
    99
  metric will decrease for each suitable edge.  A union over all edges proves
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   100
  a LEADSTO assertion that the metric decreases if we are not at a fixedpoint.
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   101
  *)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   102
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   103
lemma Compl_fixedpoint: "- fixedpoint = (\<Union>(u,v)\<in>edges. {s. s u & ~ s v})"
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13806
diff changeset
   104
apply (simp add: FP_fixedpoint [symmetric] Rprg_def mk_total_program_def)
26825
0373ed6f04b1 Manually applied subset_antisym in proof of Compl_fixedpoint, because it is
berghofe
parents: 24853
diff changeset
   105
apply (rule subset_antisym)
0373ed6f04b1 Manually applied subset_antisym in proof of Compl_fixedpoint, because it is
berghofe
parents: 24853
diff changeset
   106
apply (auto simp add: Compl_FP UN_UN_flatten del: subset_antisym)
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   107
 apply (rule fun_upd_idem, force)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   108
apply (force intro!: rev_bexI simp add: fun_upd_idem_iff)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   109
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   110
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   111
lemma Diff_fixedpoint:
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   112
     "A - fixedpoint = (\<Union>(u,v)\<in>edges. A \<inter> {s. s u & ~ s v})"
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   113
by (simp add: Diff_eq Compl_fixedpoint, blast)
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   114
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   115
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   116
(*** Progress ***)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   117
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   118
lemma Suc_metric: "~ s x ==> Suc (metric (s(x:=True))) = metric s"
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   119
apply (unfold metric_def)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   120
apply (subgoal_tac "{v. ~ (s (x:=True)) v} = {v. ~ s v} - {x}")
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   121
 prefer 2 apply force
24853
aab5798e5a33 added lemmas
nipkow
parents: 18556
diff changeset
   122
apply (simp add: card_Suc_Diff1 del:card_Diff_insert)
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   123
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   124
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   125
lemma metric_less [intro!]: "~ s x ==> metric (s(x:=True)) < metric s"
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   126
by (erule Suc_metric [THEN subst], blast)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   127
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   128
lemma metric_le: "metric (s(y:=s x | s y)) \<le> metric s"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   129
apply (case_tac "s x --> s y")
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   130
apply (auto intro: less_imp_le simp add: fun_upd_idem)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   131
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   132
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   133
lemma LeadsTo_Diff_fixedpoint:
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   134
     "Rprg \<in> ((metric-`{m}) - fixedpoint) LeadsTo (metric-`(lessThan m))"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   135
apply (simp (no_asm) add: Diff_fixedpoint Rprg_def)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   136
apply (rule LeadsTo_UN, auto)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   137
apply (ensures_tac "asgt a b")
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   138
 prefer 2 apply blast
16796
140f1e0ea846 generlization of some "nat" theorems
paulson
parents: 16417
diff changeset
   139
apply (simp (no_asm_use) add: linorder_not_less)
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   140
apply (drule metric_le [THEN order_antisym]) 
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   141
apply (auto elim: less_not_refl3 [THEN [2] rev_notE])
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   142
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   143
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   144
lemma LeadsTo_Un_fixedpoint:
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   145
     "Rprg \<in> (metric-`{m}) LeadsTo (metric-`(lessThan m) \<union> fixedpoint)"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   146
apply (rule LeadsTo_Diff [OF LeadsTo_Diff_fixedpoint [THEN LeadsTo_weaken_R]
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   147
                             subset_imp_LeadsTo], auto)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   148
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   149
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   150
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   151
(*Execution in any state leads to a fixedpoint (i.e. can terminate)*)
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   152
lemma LeadsTo_fixedpoint: "Rprg \<in> UNIV LeadsTo fixedpoint"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   153
apply (rule LessThan_induct, auto)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   154
apply (rule LeadsTo_Un_fixedpoint)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   155
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   156
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
   157
lemma LeadsTo_correct: "Rprg \<in> UNIV LeadsTo { %v. (init, v) \<in> edges^* }"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   158
apply (subst fixedpoint_invariant_correct [symmetric])
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   159
apply (rule Always_LeadsTo_weaken [OF reach_invariant LeadsTo_fixedpoint], auto)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   160
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12338
diff changeset
   161
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   162
end