| author | wenzelm | 
| Fri, 08 Apr 2016 20:52:40 +0200 | |
| changeset 62914 | 930a30c1a9af | 
| parent 61943 | 7fba644ed827 | 
| child 63146 | f1ecba0272f9 | 
| permissions | -rw-r--r-- | 
| 5899 | 1 | (* Title: HOL/UNITY/PPROD.thy | 
| 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 3 | Copyright 1998 University of Cambridge | |
| 4 | ||
| 7188 
2bc63a44721b
re-organization of theorems from Alloc and PPROD, partly into new theory
 paulson parents: 
6842diff
changeset | 5 | Abstraction over replicated components (PLam) | 
| 
2bc63a44721b
re-organization of theorems from Alloc and PPROD, partly into new theory
 paulson parents: 
6842diff
changeset | 6 | General products of programs (Pi operation) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 7 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 8 | Some dead wood here! | 
| 5899 | 9 | *) | 
| 10 | ||
| 16417 | 11 | theory PPROD imports Lift_prog begin | 
| 5899 | 12 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35054diff
changeset | 13 | definition PLam :: "[nat set, nat => ('b * ((nat=>'b) * 'c)) program]
 | 
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35054diff
changeset | 14 | => ((nat=>'b) * 'c) program" where | 
| 13805 | 15 | "PLam I F == \<Squnion>i \<in> I. lift i (F i)" | 
| 5899 | 16 | |
| 17 | syntax | |
| 35054 | 18 |   "_PLam" :: "[pttrn, nat set, 'b set] => (nat => 'b) set"  ("(3plam _:_./ _)" 10)
 | 
| 5899 | 19 | translations | 
| 35054 | 20 | "plam x : A. B" == "CONST PLam A (%x. B)" | 
| 5899 | 21 | |
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 22 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 23 | (*** Basic properties ***) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 24 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 25 | lemma Init_PLam [simp]: "Init (PLam I F) = (\<Inter>i \<in> I. lift_set i (Init (F i)))" | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 26 | by (simp add: PLam_def lift_def lift_set_def) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 27 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 28 | lemma PLam_empty [simp]: "PLam {} F = SKIP"
 | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 29 | by (simp add: PLam_def) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 30 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 31 | lemma PLam_SKIP [simp]: "(plam i : I. SKIP) = SKIP" | 
| 46577 | 32 | by (simp add: PLam_def JN_constant) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 33 | |
| 60773 | 34 | lemma PLam_insert: "PLam (insert i I) F = (lift i (F i)) \<squnion> (PLam I F)" | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 35 | by (unfold PLam_def, auto) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 36 | |
| 13805 | 37 | lemma PLam_component_iff: "((PLam I F) \<le> H) = (\<forall>i \<in> I. lift i (F i) \<le> H)" | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 38 | by (simp add: PLam_def JN_component_iff) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 39 | |
| 13805 | 40 | lemma component_PLam: "i \<in> I ==> lift i (F i) \<le> (PLam I F)" | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 41 | apply (unfold PLam_def) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 42 | (*blast_tac doesn't use HO unification*) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 43 | apply (fast intro: component_JN) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 44 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 45 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 46 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 47 | (** Safety & Progress: but are they used anywhere? **) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 48 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 49 | lemma PLam_constrains: | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 50 | "[| i \<in> I; \<forall>j. F j \<in> preserves snd |] | 
| 61943 | 51 | ==> (PLam I F \<in> (lift_set i (A \<times> UNIV)) co | 
| 52 | (lift_set i (B \<times> UNIV))) = | |
| 53 | (F i \<in> (A \<times> UNIV) co (B \<times> UNIV))" | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 54 | apply (simp add: PLam_def JN_constrains) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 55 | apply (subst insert_Diff [symmetric], assumption) | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 56 | apply (simp add: lift_constrains) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 57 | apply (blast intro: constrains_imp_lift_constrains) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 58 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 59 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 60 | lemma PLam_stable: | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 61 | "[| i \<in> I; \<forall>j. F j \<in> preserves snd |] | 
| 61943 | 62 | ==> (PLam I F \<in> stable (lift_set i (A \<times> UNIV))) = | 
| 63 | (F i \<in> stable (A \<times> UNIV))" | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 64 | by (simp add: stable_def PLam_constrains) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 65 | |
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 66 | lemma PLam_transient: | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 67 | "i \<in> I ==> | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 68 | PLam I F \<in> transient A = (\<exists>i \<in> I. lift i (F i) \<in> transient A)" | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 69 | by (simp add: JN_transient PLam_def) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 70 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 71 | text{*This holds because the @{term "F j"} cannot change @{term "lift_set i"}*}
 | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 72 | lemma PLam_ensures: | 
| 61943 | 73 | "[| i \<in> I; F i \<in> (A \<times> UNIV) ensures (B \<times> UNIV); | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 74 | \<forall>j. F j \<in> preserves snd |] | 
| 61943 | 75 | ==> PLam I F \<in> lift_set i (A \<times> UNIV) ensures lift_set i (B \<times> UNIV)" | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 76 | apply (simp add: ensures_def PLam_constrains PLam_transient | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 77 | lift_set_Un_distrib [symmetric] lift_set_Diff_distrib [symmetric] | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 78 | Times_Un_distrib1 [symmetric] Times_Diff_distrib1 [symmetric]) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 79 | apply (rule rev_bexI, assumption) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 80 | apply (simp add: lift_transient) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 81 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 82 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 83 | lemma PLam_leadsTo_Basis: | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 84 | "[| i \<in> I; | 
| 61943 | 85 | F i \<in> ((A \<times> UNIV) - (B \<times> UNIV)) co | 
| 86 | ((A \<times> UNIV) \<union> (B \<times> UNIV)); | |
| 87 | F i \<in> transient ((A \<times> UNIV) - (B \<times> UNIV)); | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 88 | \<forall>j. F j \<in> preserves snd |] | 
| 61943 | 89 | ==> PLam I F \<in> lift_set i (A \<times> UNIV) leadsTo lift_set i (B \<times> UNIV)" | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 90 | by (rule PLam_ensures [THEN leadsTo_Basis], rule_tac [2] ensuresI) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 91 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 92 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 93 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 94 | (** invariant **) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 95 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 96 | lemma invariant_imp_PLam_invariant: | 
| 61943 | 97 | "[| F i \<in> invariant (A \<times> UNIV); i \<in> I; | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 98 | \<forall>j. F j \<in> preserves snd |] | 
| 61943 | 99 | ==> PLam I F \<in> invariant (lift_set i (A \<times> UNIV))" | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 100 | by (auto simp add: PLam_stable invariant_def) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 101 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 102 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 103 | lemma PLam_preserves_fst [simp]: | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 104 | "\<forall>j. F j \<in> preserves snd | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 105 | ==> (PLam I F \<in> preserves (v o sub j o fst)) = | 
| 13805 | 106 | (if j \<in> I then F j \<in> preserves (v o fst) else True)" | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 107 | by (simp add: PLam_def lift_preserves_sub) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 108 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 109 | lemma PLam_preserves_snd [simp,intro]: | 
| 13805 | 110 | "\<forall>j. F j \<in> preserves snd ==> PLam I F \<in> preserves snd" | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 111 | by (simp add: PLam_def lift_preserves_snd_I) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 112 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 113 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 114 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 115 | (*** guarantees properties ***) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 116 | |
| 14150 | 117 | text{*This rule looks unsatisfactory because it refers to @{term lift}. 
 | 
| 118 | One must use | |
| 119 |   @{text lift_guarantees_eq_lift_inv} to rewrite the first subgoal and
 | |
| 120 |   something like @{text lift_preserves_sub} to rewrite the third.  However
 | |
| 121 | there's no obvious alternative for the third premise.*} | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 122 | lemma guarantees_PLam_I: | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 123 | "[| lift i (F i): X guarantees Y; i \<in> I; | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 124 | OK I (%i. lift i (F i)) |] | 
| 13805 | 125 | ==> (PLam I F) \<in> X guarantees Y" | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 126 | apply (unfold PLam_def) | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 127 | apply (simp add: guarantees_JN_I) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 128 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 129 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 130 | lemma Allowed_PLam [simp]: | 
| 13805 | 131 | "Allowed (PLam I F) = (\<Inter>i \<in> I. lift i ` Allowed(F i))" | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 132 | by (simp add: PLam_def) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 133 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 134 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 135 | lemma PLam_preserves [simp]: | 
| 13805 | 136 | "(PLam I F) \<in> preserves v = (\<forall>i \<in> I. F i \<in> preserves (v o lift_map i))" | 
| 137 | by (simp add: PLam_def lift_def rename_preserves) | |
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 138 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 139 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 140 | (**UNUSED | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 141 | (*The f0 premise ensures that the product is well-defined.*) | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 142 | lemma PLam_invariant_imp_invariant: | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 143 | "[| PLam I F \<in> invariant (lift_set i A); i \<in> I; | 
| 13805 | 144 | f0: Init (PLam I F) |] ==> F i \<in> invariant A" | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 145 | apply (auto simp add: invariant_def) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 146 | apply (drule_tac c = "f0 (i:=x) " in subsetD) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 147 | apply auto | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 148 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 149 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 150 | lemma PLam_invariant: | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 151 | "[| i \<in> I; f0: Init (PLam I F) |] | 
| 13805 | 152 | ==> (PLam I F \<in> invariant (lift_set i A)) = (F i \<in> invariant A)" | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 153 | apply (blast intro: invariant_imp_PLam_invariant PLam_invariant_imp_invariant) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 154 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 155 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 156 | (*The f0 premise isn't needed if F is a constant program because then | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 157 | we get an initial state by replicating that of F*) | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 158 | lemma reachable_PLam: | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 159 | "i \<in> I | 
| 13805 | 160 | ==> ((plam x \<in> I. F) \<in> invariant (lift_set i A)) = (F \<in> invariant A)" | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 161 | apply (auto simp add: invariant_def) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 162 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 163 | **) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 164 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 165 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 166 | (**UNUSED | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 167 | (** Reachability **) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 168 | |
| 13805 | 169 | Goal "[| f \<in> reachable (PLam I F); i \<in> I |] ==> f i \<in> reachable (F i)" | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 170 | apply (erule reachable.induct) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 171 | apply (auto intro: reachable.intrs) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 172 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 173 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 174 | (*Result to justify a re-organization of this file*) | 
| 13805 | 175 |     lemma "{f. \<forall>i \<in> I. f i \<in> R i} = (\<Inter>i \<in> I. lift_set i (R i))"
 | 
| 13798 | 176 | by auto | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 177 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 178 | lemma reachable_PLam_subset1: | 
| 13805 | 179 | "reachable (PLam I F) \<subseteq> (\<Inter>i \<in> I. lift_set i (reachable (F i)))" | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 180 | apply (force dest!: reachable_PLam) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 181 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 182 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 183 | (*simplify using reachable_lift??*) | 
| 13798 | 184 | lemma reachable_lift_Join_PLam [rule_format]: | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 185 | "[| i \<notin> I; A \<in> reachable (F i) |] | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 186 | ==> \<forall>f. f \<in> reachable (PLam I F) | 
| 60773 | 187 | --> f(i:=A) \<in> reachable (lift i (F i) \<squnion> PLam I F)" | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 188 | apply (erule reachable.induct) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 189 | apply (ALLGOALS Clarify_tac) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 190 | apply (erule reachable.induct) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 191 | (*Init, Init case*) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 192 | apply (force intro: reachable.intrs) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 193 | (*Init of F, action of PLam F case*) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 194 | apply (rule_tac act = act in reachable.Acts) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 195 | apply force | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 196 | apply assumption | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 197 | apply (force intro: ext) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 198 | (*induction over the 2nd "reachable" assumption*) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 199 | apply (erule_tac xa = f in reachable.induct) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 200 | (*Init of PLam F, action of F case*) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 201 | apply (rule_tac act = "lift_act i act" in reachable.Acts) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 202 | apply force | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 203 | apply (force intro: reachable.Init) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 204 | apply (force intro: ext simp add: lift_act_def) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 205 | (*last case: an action of PLam I F*) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 206 | apply (rule_tac act = acta in reachable.Acts) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 207 | apply force | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 208 | apply assumption | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 209 | apply (force intro: ext) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 210 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 211 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 212 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 213 | (*The index set must be finite: otherwise infinitely many copies of F can | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 214 | perform actions, and PLam can never catch up in finite time.*) | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 215 | lemma reachable_PLam_subset2: | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 216 | "finite I | 
| 13805 | 217 | ==> (\<Inter>i \<in> I. lift_set i (reachable (F i))) \<subseteq> reachable (PLam I F)" | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 218 | apply (erule finite_induct) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 219 | apply (simp (no_asm)) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 220 | apply (force dest: reachable_lift_Join_PLam simp add: PLam_insert) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 221 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 222 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 223 | lemma reachable_PLam_eq: | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 224 | "finite I ==> | 
| 13805 | 225 | reachable (PLam I F) = (\<Inter>i \<in> I. lift_set i (reachable (F i)))" | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 226 | apply (REPEAT_FIRST (ares_tac [equalityI, reachable_PLam_subset1, reachable_PLam_subset2])) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 227 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 228 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 229 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 230 | (** Co **) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 231 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 232 | lemma Constrains_imp_PLam_Constrains: | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 233 | "[| F i \<in> A Co B; i \<in> I; finite I |] | 
| 13805 | 234 | ==> PLam I F \<in> (lift_set i A) Co (lift_set i B)" | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 235 | apply (auto simp add: Constrains_def Collect_conj_eq [symmetric] reachable_PLam_eq) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 236 | apply (auto simp add: constrains_def PLam_def) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 237 | apply (REPEAT (blast intro: reachable.intrs)) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 238 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 239 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 240 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 241 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 242 | lemma PLam_Constrains: | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 243 | "[| i \<in> I; finite I; f0: Init (PLam I F) |] | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 244 | ==> (PLam I F \<in> (lift_set i A) Co (lift_set i B)) = | 
| 13805 | 245 | (F i \<in> A Co B)" | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 246 | apply (blast intro: Constrains_imp_PLam_Constrains PLam_Constrains_imp_Constrains) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 247 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 248 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 249 | lemma PLam_Stable: | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 250 | "[| i \<in> I; finite I; f0: Init (PLam I F) |] | 
| 13805 | 251 | ==> (PLam I F \<in> Stable (lift_set i A)) = (F i \<in> Stable A)" | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 252 | apply (simp del: Init_PLam add: Stable_def PLam_Constrains) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 253 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 254 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 255 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 256 | (** const_PLam (no dependence on i) doesn't require the f0 premise **) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 257 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 258 | lemma const_PLam_Constrains: | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 259 | "[| i \<in> I; finite I |] | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 260 | ==> ((plam x \<in> I. F) \<in> (lift_set i A) Co (lift_set i B)) = | 
| 13805 | 261 | (F \<in> A Co B)" | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 262 | apply (blast intro: Constrains_imp_PLam_Constrains const_PLam_Constrains_imp_Constrains) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 263 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 264 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 265 | lemma const_PLam_Stable: | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 266 | "[| i \<in> I; finite I |] | 
| 13805 | 267 | ==> ((plam x \<in> I. F) \<in> Stable (lift_set i A)) = (F \<in> Stable A)" | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 268 | apply (simp add: Stable_def const_PLam_Constrains) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 269 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 270 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 271 | lemma const_PLam_Increasing: | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
16417diff
changeset | 272 | "[| i \<in> I; finite I |] | 
| 13805 | 273 | ==> ((plam x \<in> I. F) \<in> Increasing (f o sub i)) = (F \<in> Increasing f)" | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 274 | apply (unfold Increasing_def) | 
| 13805 | 275 |     apply (subgoal_tac "\<forall>z. {s. z \<subseteq> (f o sub i) s} = lift_set i {s. z \<subseteq> f s}")
 | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 276 | apply (asm_simp_tac (simpset () add: lift_set_sub) 2) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 277 | apply (simp add: finite_lessThan const_PLam_Stable) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 278 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 279 | **) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 280 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
8251diff
changeset | 281 | |
| 5899 | 282 | end |