| author | wenzelm | 
| Tue, 20 Sep 2005 14:13:20 +0200 | |
| changeset 17517 | 9dc9d3005ed2 | 
| parent 16796 | 140f1e0ea846 | 
| child 18556 | dc39832e9280 | 
| 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/Reach.thy | 
| 
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: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 4 | Copyright 1998 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 Directed Graphs. From Chandy and Misra, section 6.4. | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 7 | [this example took only four days!] | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 8 | *) | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 9 | |
| 16417 | 10 | theory Reach imports UNITY_Main begin | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 11 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 12 | typedecl vertex | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 13 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 14 | types state = "vertex=>bool" | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 15 | |
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 16 | consts | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 17 | init :: "vertex" | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 18 | |
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 19 | edges :: "(vertex*vertex) set" | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 20 | |
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 21 | constdefs | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 22 | |
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 23 | asgt :: "[vertex,vertex] => (state*state) set" | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 24 |     "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 | 25 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 26 | Rprg :: "state program" | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13806diff
changeset | 27 |     "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 | 28 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 29 | reach_invariant :: "state set" | 
| 13806 | 30 |     "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 | 31 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 32 | fixedpoint :: "state set" | 
| 13806 | 33 |     "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 | 34 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 35 | metric :: "state => nat" | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 36 |     "metric s == card {v. ~ s v}"
 | 
| 
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: 
12338diff
changeset | 38 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 39 | text{**We assume that the set of vertices is finite*}
 | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 40 | axioms | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 41 | finite_graph: "finite (UNIV :: vertex set)" | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 42 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 43 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 44 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 45 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 46 | (*TO SIMPDATA.ML?? FOR CLASET?? *) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 47 | lemma ifE [elim!]: | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 48 | "[| if P then Q else R; | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 49 | [| P; Q |] ==> S; | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 50 | [| ~ P; R |] ==> S |] ==> S" | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 51 | by (simp split: split_if_asm) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 52 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 53 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 54 | declare Rprg_def [THEN def_prg_Init, simp] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 55 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 56 | declare asgt_def [THEN def_act_simp,simp] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 57 | |
| 14150 | 58 | text{*All vertex sets are finite*}
 | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 59 | declare finite_subset [OF subset_UNIV finite_graph, iff] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 60 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 61 | declare reach_invariant_def [THEN def_set_simp, simp] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 62 | |
| 13806 | 63 | lemma reach_invariant: "Rprg \<in> Always reach_invariant" | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 64 | apply (rule AlwaysI, force) | 
| 16184 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
15097diff
changeset | 65 | apply (unfold Rprg_def, safety) | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 66 | apply (blast intro: rtrancl_trans) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 67 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 68 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 69 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 70 | (*** Fixedpoint ***) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 71 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 72 | (*If it reaches a fixedpoint, it has found a solution*) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 73 | lemma fixedpoint_invariant_correct: | 
| 13806 | 74 |      "fixedpoint \<inter> reach_invariant = { %v. (init, v) \<in> edges^* }"
 | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 75 | apply (unfold fixedpoint_def) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 76 | apply (rule equalityI) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 77 | apply (auto intro!: ext) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 78 | apply (erule rtrancl_induct, auto) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 79 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 80 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 81 | lemma lemma1: | 
| 13806 | 82 | "FP Rprg \<subseteq> fixedpoint" | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13806diff
changeset | 83 | apply (simp add: FP_def fixedpoint_def Rprg_def mk_total_program_def) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13806diff
changeset | 84 | apply (auto simp add: stable_def constrains_def) | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 85 | apply (drule bspec, assumption) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 86 | apply (simp add: Image_singleton image_iff) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 87 | apply (drule fun_cong, auto) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 88 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 89 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 90 | lemma lemma2: | 
| 13806 | 91 | "fixedpoint \<subseteq> FP Rprg" | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13806diff
changeset | 92 | apply (simp add: FP_def fixedpoint_def Rprg_def mk_total_program_def) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13806diff
changeset | 93 | apply (auto intro!: ext simp add: stable_def constrains_def) | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 94 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 95 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 96 | lemma FP_fixedpoint: "FP Rprg = fixedpoint" | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 97 | by (rule equalityI [OF lemma1 lemma2]) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 98 | |
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 99 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 100 | (*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: 
12338diff
changeset | 101 | 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: 
12338diff
changeset | 102 | 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: 
12338diff
changeset | 103 | 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: 
12338diff
changeset | 104 | *) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 105 | |
| 13806 | 106 | 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: 
13806diff
changeset | 107 | apply (simp add: FP_fixedpoint [symmetric] Rprg_def mk_total_program_def) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13806diff
changeset | 108 | apply (auto simp add: Compl_FP UN_UN_flatten) | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 109 | apply (rule fun_upd_idem, force) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 110 | apply (force intro!: rev_bexI simp add: fun_upd_idem_iff) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 111 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 112 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 113 | lemma Diff_fixedpoint: | 
| 13806 | 114 |      "A - fixedpoint = (\<Union>(u,v)\<in>edges. A \<inter> {s. s u & ~ s v})"
 | 
| 115 | by (simp add: Diff_eq Compl_fixedpoint, blast) | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 116 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 117 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 118 | (*** Progress ***) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 119 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 120 | lemma Suc_metric: "~ s x ==> Suc (metric (s(x:=True))) = metric s" | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 121 | apply (unfold metric_def) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 122 | apply (subgoal_tac "{v. ~ (s (x:=True)) v} = {v. ~ s v} - {x}")
 | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 123 | prefer 2 apply force | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 124 | apply (simp add: card_Suc_Diff1) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 125 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 126 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 127 | lemma metric_less [intro!]: "~ s x ==> metric (s(x:=True)) < metric s" | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 128 | by (erule Suc_metric [THEN subst], blast) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 129 | |
| 13806 | 130 | 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: 
12338diff
changeset | 131 | apply (case_tac "s x --> s y") | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 132 | apply (auto intro: less_imp_le simp add: fun_upd_idem) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 133 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 134 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 135 | lemma LeadsTo_Diff_fixedpoint: | 
| 13806 | 136 |      "Rprg \<in> ((metric-`{m}) - fixedpoint) LeadsTo (metric-`(lessThan m))"
 | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 137 | apply (simp (no_asm) add: Diff_fixedpoint Rprg_def) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 138 | apply (rule LeadsTo_UN, auto) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 139 | apply (ensures_tac "asgt a b") | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 140 | prefer 2 apply blast | 
| 16796 | 141 | apply (simp (no_asm_use) add: linorder_not_less) | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 142 | apply (drule metric_le [THEN order_antisym]) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 143 | apply (auto elim: less_not_refl3 [THEN [2] rev_notE]) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 144 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 145 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 146 | lemma LeadsTo_Un_fixedpoint: | 
| 13806 | 147 |      "Rprg \<in> (metric-`{m}) LeadsTo (metric-`(lessThan m) \<union> fixedpoint)"
 | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 148 | apply (rule LeadsTo_Diff [OF LeadsTo_Diff_fixedpoint [THEN LeadsTo_weaken_R] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 149 | subset_imp_LeadsTo], auto) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 150 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 151 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 152 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 153 | (*Execution in any state leads to a fixedpoint (i.e. can terminate)*) | 
| 13806 | 154 | lemma LeadsTo_fixedpoint: "Rprg \<in> UNIV LeadsTo fixedpoint" | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 155 | apply (rule LessThan_induct, auto) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 156 | apply (rule LeadsTo_Un_fixedpoint) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 157 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 158 | |
| 13806 | 159 | 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: 
12338diff
changeset | 160 | apply (subst fixedpoint_invariant_correct [symmetric]) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 161 | apply (rule Always_LeadsTo_weaken [OF reach_invariant LeadsTo_fixedpoint], auto) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 162 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
12338diff
changeset | 163 | |
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 164 | end |