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