| author | isatest | 
| Thu, 22 Sep 2005 00:30:31 +0200 | |
| changeset 17578 | e07af5fad73f | 
| parent 16417 | 9bc16273c2d4 | 
| child 24892 | c663e675e177 | 
| permissions | -rw-r--r-- | 
| 14052 | 1 | (* Title: ZF/UNITY/Increasing | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 2 | ID: $Id \<in> Increasing.thy,v 1.2 2003/06/02 09:17:52 paulson Exp $ | 
| 14052 | 3 | Author: Sidi O Ehmety, Cambridge University Computer Laboratory | 
| 4 | Copyright 2001 University of Cambridge | |
| 5 | ||
| 6 | Increasing's parameters are a state function f, a domain A and an order | |
| 7 | relation r over the domain A. | |
| 8 | *) | |
| 9 | ||
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 10 | header{*Charpentier's "Increasing" Relation*}
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 11 | |
| 16417 | 12 | theory Increasing imports Constrains Monotonicity begin | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 13 | |
| 14052 | 14 | constdefs | 
| 15 | ||
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 16 |   increasing :: "[i, i, i=>i] => i" ("increasing[_]'(_, _')")
 | 
| 14052 | 17 | "increasing[A](r, f) == | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 18 |     {F \<in> program. (\<forall>k \<in> A. F \<in> stable({s \<in> state. <k, f(s)> \<in> r})) &
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 19 | (\<forall>x \<in> state. f(x):A)}" | 
| 14052 | 20 | |
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 21 |   Increasing :: "[i, i, i=>i] => i" ("Increasing[_]'(_, _')")
 | 
| 14052 | 22 | "Increasing[A](r, f) == | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 23 |     {F \<in> program. (\<forall>k \<in> A. F \<in> Stable({s \<in> state. <k, f(s)> \<in> r})) &
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 24 | (\<forall>x \<in> state. f(x):A)}" | 
| 14052 | 25 | |
| 26 | syntax | |
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 27 |   IncWrt ::  "[i=>i, i, i] => i" ("(_ IncreasingWrt _ '/ _)" [60, 0, 60] 60)
 | 
| 14052 | 28 | |
| 29 | translations | |
| 30 | "IncWrt(f,r,A)" => "Increasing[A](r,f)" | |
| 31 | ||
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 32 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 33 | (** increasing **) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 34 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 35 | lemma increasing_type: "increasing[A](r, f) <= program" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 36 | by (unfold increasing_def, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 37 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 38 | lemma increasing_into_program: "F \<in> increasing[A](r, f) ==> F \<in> program" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 39 | by (unfold increasing_def, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 40 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 41 | lemma increasing_imp_stable: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 42 | "[| F \<in> increasing[A](r, f); x \<in> A |] ==>F \<in> stable({s \<in> state. <x, f(s)>:r})"
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 43 | by (unfold increasing_def, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 44 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 45 | lemma increasingD: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 46 | "F \<in> increasing[A](r,f) ==> F \<in> program & (\<exists>a. a \<in> A) & (\<forall>s \<in> state. f(s):A)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 47 | apply (unfold increasing_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 48 | apply (subgoal_tac "\<exists>x. x \<in> state") | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 49 | apply (auto dest: stable_type [THEN subsetD] intro: st0_in_state) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 50 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 51 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 52 | lemma increasing_constant [simp]: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 53 | "F \<in> increasing[A](r, %s. c) <-> F \<in> program & c \<in> A" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 54 | apply (unfold increasing_def stable_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 55 | apply (subgoal_tac "\<exists>x. x \<in> state") | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 56 | apply (auto dest: stable_type [THEN subsetD] intro: st0_in_state) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 57 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 58 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 59 | lemma subset_increasing_comp: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 60 | "[| mono1(A, r, B, s, g); refl(A, r); trans[B](s) |] ==> | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 61 | increasing[A](r, f) <= increasing[B](s, g comp f)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 62 | apply (unfold increasing_def stable_def part_order_def | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 63 | constrains_def mono1_def metacomp_def, clarify, simp) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 64 | apply clarify | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 65 | apply (subgoal_tac "xa \<in> state") | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 66 | prefer 2 apply (blast dest!: ActsD) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 67 | apply (subgoal_tac "<f (xb), f (xb) >:r") | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 68 | prefer 2 apply (force simp add: refl_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 69 | apply (rotate_tac 5) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 70 | apply (drule_tac x = "f (xb) " in bspec) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 71 | apply (rotate_tac [2] -1) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 72 | apply (drule_tac [2] x = act in bspec, simp_all) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 73 | apply (drule_tac A = "act``?u" and c = xa in subsetD, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 74 | apply (drule_tac x = "f (xa) " and x1 = "f (xb) " in bspec [THEN bspec]) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 75 | apply (rule_tac [3] b = "g (f (xb))" and A = B in trans_onD) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 76 | apply simp_all | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 77 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 78 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 79 | lemma imp_increasing_comp: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 80 | "[| F \<in> increasing[A](r, f); mono1(A, r, B, s, g); | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 81 | refl(A, r); trans[B](s) |] ==> F \<in> increasing[B](s, g comp f)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 82 | by (rule subset_increasing_comp [THEN subsetD], auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 83 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 84 | lemma strict_increasing: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 85 | "increasing[nat](Le, f) <= increasing[nat](Lt, f)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 86 | by (unfold increasing_def Lt_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 87 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 88 | lemma strict_gt_increasing: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 89 | "increasing[nat](Ge, f) <= increasing[nat](Gt, f)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 90 | apply (unfold increasing_def Gt_def Ge_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 91 | apply (erule natE) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 92 | apply (auto simp add: stable_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 93 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 94 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 95 | (** Increasing **) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 96 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 97 | lemma increasing_imp_Increasing: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 98 | "F \<in> increasing[A](r, f) ==> F \<in> Increasing[A](r, f)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 99 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 100 | apply (unfold increasing_def Increasing_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 101 | apply (auto intro: stable_imp_Stable) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 102 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 103 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 104 | lemma Increasing_type: "Increasing[A](r, f) <= program" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 105 | by (unfold Increasing_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 106 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 107 | lemma Increasing_into_program: "F \<in> Increasing[A](r, f) ==> F \<in> program" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 108 | by (unfold Increasing_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 109 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 110 | lemma Increasing_imp_Stable: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 111 | "[| F \<in> Increasing[A](r, f); a \<in> A |] ==> F \<in> Stable({s \<in> state. <a,f(s)>:r})"
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 112 | by (unfold Increasing_def, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 113 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 114 | lemma IncreasingD: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 115 | "F \<in> Increasing[A](r, f) ==> F \<in> program & (\<exists>a. a \<in> A) & (\<forall>s \<in> state. f(s):A)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 116 | apply (unfold Increasing_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 117 | apply (subgoal_tac "\<exists>x. x \<in> state") | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 118 | apply (auto intro: st0_in_state) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 119 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 120 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 121 | lemma Increasing_constant [simp]: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 122 | "F \<in> Increasing[A](r, %s. c) <-> F \<in> program & (c \<in> A)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 123 | apply (subgoal_tac "\<exists>x. x \<in> state") | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 124 | apply (auto dest!: IncreasingD intro: st0_in_state increasing_imp_Increasing) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 125 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 126 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 127 | lemma subset_Increasing_comp: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 128 | "[| mono1(A, r, B, s, g); refl(A, r); trans[B](s) |] ==> | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 129 | Increasing[A](r, f) <= Increasing[B](s, g comp f)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 130 | apply (unfold Increasing_def Stable_def Constrains_def part_order_def | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 131 | constrains_def mono1_def metacomp_def, safe) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 132 | apply (simp_all add: ActsD) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 133 | apply (subgoal_tac "xb \<in> state & xa \<in> state") | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 134 | prefer 2 apply (simp add: ActsD) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 135 | apply (subgoal_tac "<f (xb), f (xb) >:r") | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 136 | prefer 2 apply (force simp add: refl_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 137 | apply (rotate_tac 5) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 138 | apply (drule_tac x = "f (xb) " in bspec) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 139 | apply simp_all | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 140 | apply clarify | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 141 | apply (rotate_tac -2) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 142 | apply (drule_tac x = act in bspec) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 143 | apply (drule_tac [2] A = "act``?u" and c = xa in subsetD, simp_all, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 144 | apply (drule_tac x = "f (xa) " and x1 = "f (xb) " in bspec [THEN bspec]) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 145 | apply (rule_tac [3] b = "g (f (xb))" and A = B in trans_onD) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 146 | apply simp_all | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 147 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 148 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 149 | lemma imp_Increasing_comp: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 150 | "[| F \<in> Increasing[A](r, f); mono1(A, r, B, s, g); refl(A, r); trans[B](s) |] | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 151 | ==> F \<in> Increasing[B](s, g comp f)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 152 | apply (rule subset_Increasing_comp [THEN subsetD], auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 153 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 154 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 155 | lemma strict_Increasing: "Increasing[nat](Le, f) <= Increasing[nat](Lt, f)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 156 | by (unfold Increasing_def Lt_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 157 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 158 | lemma strict_gt_Increasing: "Increasing[nat](Ge, f)<= Increasing[nat](Gt, f)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 159 | apply (unfold Increasing_def Ge_def Gt_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 160 | apply (erule natE) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 161 | apply (auto simp add: Stable_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 162 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 163 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 164 | (** Two-place monotone operations **) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 165 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 166 | lemma imp_increasing_comp2: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 167 | "[| F \<in> increasing[A](r, f); F \<in> increasing[B](s, g); | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 168 | mono2(A, r, B, s, C, t, h); refl(A, r); refl(B, s); trans[C](t) |] | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 169 | ==> F \<in> increasing[C](t, %x. h(f(x), g(x)))" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 170 | apply (unfold increasing_def stable_def | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 171 | part_order_def constrains_def mono2_def, clarify, simp) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 172 | apply clarify | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 173 | apply (rename_tac xa xb) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 174 | apply (subgoal_tac "xb \<in> state & xa \<in> state") | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 175 | prefer 2 apply (blast dest!: ActsD) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 176 | apply (subgoal_tac "<f (xb), f (xb) >:r & <g (xb), g (xb) >:s") | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 177 | prefer 2 apply (force simp add: refl_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 178 | apply (rotate_tac 6) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 179 | apply (drule_tac x = "f (xb) " in bspec) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 180 | apply (rotate_tac [2] 1) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 181 | apply (drule_tac [2] x = "g (xb) " in bspec) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 182 | apply simp_all | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 183 | apply (rotate_tac -1) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 184 | apply (drule_tac x = act in bspec) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 185 | apply (rotate_tac [2] -3) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 186 | apply (drule_tac [2] x = act in bspec, simp_all) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 187 | apply (drule_tac A = "act``?u" and c = xa in subsetD) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 188 | apply (drule_tac [2] A = "act``?u" and c = xa in subsetD, blast, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 189 | apply (rotate_tac -4) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 190 | apply (drule_tac x = "f (xa) " and x1 = "f (xb) " in bspec [THEN bspec]) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 191 | apply (rotate_tac [3] -1) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 192 | apply (drule_tac [3] x = "g (xa) " and x1 = "g (xb) " in bspec [THEN bspec]) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 193 | apply simp_all | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 194 | apply (rule_tac b = "h (f (xb), g (xb))" and A = C in trans_onD) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 195 | apply simp_all | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 196 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 197 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 198 | lemma imp_Increasing_comp2: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 199 | "[| F \<in> Increasing[A](r, f); F \<in> Increasing[B](s, g); | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 200 | mono2(A, r, B, s, C, t, h); refl(A, r); refl(B, s); trans[C](t) |] ==> | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 201 | F \<in> Increasing[C](t, %x. h(f(x), g(x)))" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 202 | apply (unfold Increasing_def stable_def | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 203 | part_order_def constrains_def mono2_def Stable_def Constrains_def, safe) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 204 | apply (simp_all add: ActsD) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 205 | apply (subgoal_tac "xa \<in> state & x \<in> state") | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 206 | prefer 2 apply (blast dest!: ActsD) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 207 | apply (subgoal_tac "<f (xa), f (xa) >:r & <g (xa), g (xa) >:s") | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 208 | prefer 2 apply (force simp add: refl_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 209 | apply (rotate_tac 6) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 210 | apply (drule_tac x = "f (xa) " in bspec) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 211 | apply (rotate_tac [2] 1) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 212 | apply (drule_tac [2] x = "g (xa) " in bspec) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 213 | apply simp_all | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 214 | apply clarify | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 215 | apply (rotate_tac -2) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 216 | apply (drule_tac x = act in bspec) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 217 | apply (rotate_tac [2] -3) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 218 | apply (drule_tac [2] x = act in bspec, simp_all) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 219 | apply (drule_tac A = "act``?u" and c = x in subsetD) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 220 | apply (drule_tac [2] A = "act``?u" and c = x in subsetD, blast, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 221 | apply (rotate_tac -9) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 222 | apply (drule_tac x = "f (x) " and x1 = "f (xa) " in bspec [THEN bspec]) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 223 | apply (rotate_tac [3] -1) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 224 | apply (drule_tac [3] x = "g (x) " and x1 = "g (xa) " in bspec [THEN bspec]) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 225 | apply simp_all | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 226 | apply (rule_tac b = "h (f (xa), g (xa))" and A = C in trans_onD) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 227 | apply simp_all | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 228 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14055diff
changeset | 229 | |
| 14052 | 230 | |
| 231 | end |