| author | Fabian Huch <huch@in.tum.de> |
| Thu, 24 Oct 2024 14:07:13 +0200 | |
| changeset 81364 | 84e4388f8ab1 |
| parent 81125 | ec121999a9cb |
| 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 |
| 81125 | 14 |
increasing :: "[i, i, i\<Rightarrow>i] \<Rightarrow> i" (\<open>(\<open>open_block notation=\<open>mixfix increasing\<close>\<close>increasing[_]'(_, _'))\<close>) |
15 |
where |
|
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
16 |
"increasing[A](r, f) \<equiv> |
| 76214 | 17 |
{F \<in> program. (\<forall>k \<in> A. F \<in> stable({s \<in> state. <k, f(s)> \<in> r})) \<and>
|
|
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
18 |
(\<forall>x \<in> state. f(x):A)}" |
| 14052 | 19 |
|
| 24893 | 20 |
definition |
| 81125 | 21 |
Increasing :: "[i, i, i\<Rightarrow>i] \<Rightarrow> i" (\<open>(\<open>open_block notation=\<open>mixfix Increasing\<close>\<close>Increasing[_]'(_, _'))\<close>) |
22 |
where |
|
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
23 |
"Increasing[A](r, f) \<equiv> |
| 76214 | 24 |
{F \<in> program. (\<forall>k \<in> A. F \<in> Stable({s \<in> state. <k, f(s)> \<in> r})) \<and>
|
|
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
25 |
(\<forall>x \<in> state. f(x):A)}" |
| 14052 | 26 |
|
| 24892 | 27 |
abbreviation (input) |
| 80917 | 28 |
IncWrt :: "[i\<Rightarrow>i, i, i] \<Rightarrow> i" (\<open>(\<open>notation=\<open>mixfix IncreasingWrt\<close>\<close>_ IncreasingWrt _ '/ _)\<close> [60, 0, 60] 60) where |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
29 |
"f IncreasingWrt r/A \<equiv> Increasing[A](r,f)" |
| 14052 | 30 |
|
|
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
31 |
|
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
32 |
(** increasing **) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
33 |
|
| 46823 | 34 |
lemma increasing_type: "increasing[A](r, f) \<subseteq> program" |
|
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
35 |
by (unfold increasing_def, blast) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
36 |
|
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
37 |
lemma increasing_into_program: "F \<in> increasing[A](r, f) \<Longrightarrow> F \<in> program" |
|
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
38 |
by (unfold increasing_def, blast) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
39 |
|
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
40 |
lemma increasing_imp_stable: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
41 |
"\<lbrakk>F \<in> increasing[A](r, f); x \<in> A\<rbrakk> \<Longrightarrow>F \<in> stable({s \<in> state. <x, f(s)>:r})"
|
|
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
42 |
by (unfold increasing_def, blast) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
43 |
|
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
44 |
lemma increasingD: |
| 76214 | 45 |
"F \<in> increasing[A](r,f) \<Longrightarrow> F \<in> program \<and> (\<exists>a. a \<in> A) \<and> (\<forall>s \<in> state. f(s):A)" |
|
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
46 |
unfolding increasing_def |
|
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
47 |
apply (subgoal_tac "\<exists>x. x \<in> state") |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
48 |
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
|
49 |
done |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
50 |
|
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
51 |
lemma increasing_constant [simp]: |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
52 |
"F \<in> increasing[A](r, \<lambda>s. c) \<longleftrightarrow> F \<in> program \<and> c \<in> A" |
| 76217 | 53 |
unfolding increasing_def stable_def |
|
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
54 |
apply (subgoal_tac "\<exists>x. x \<in> state") |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
55 |
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
|
56 |
done |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
57 |
|
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
58 |
lemma subset_increasing_comp: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
59 |
"\<lbrakk>mono1(A, r, B, s, g); refl(A, r); trans[B](s)\<rbrakk> \<Longrightarrow> |
| 46823 | 60 |
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
|
61 |
apply (unfold increasing_def stable_def part_order_def |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
62 |
constrains_def mono1_def metacomp_def, clarify, simp) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
63 |
apply clarify |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
64 |
apply (subgoal_tac "xa \<in> state") |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
65 |
prefer 2 apply (blast dest!: ActsD) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
66 |
apply (subgoal_tac "<f (xb), f (xb) >:r") |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
67 |
prefer 2 apply (force simp add: refl_def) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
68 |
apply (rotate_tac 5) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
69 |
apply (drule_tac x = "f (xb) " in bspec) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
70 |
apply (rotate_tac [2] -1) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
71 |
apply (drule_tac [2] x = act in bspec, simp_all) |
| 59788 | 72 |
apply (drule_tac A = "act``u" and c = xa for u in subsetD, blast) |
73 |
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
|
74 |
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
|
75 |
apply simp_all |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
76 |
done |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
77 |
|
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
78 |
lemma imp_increasing_comp: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
79 |
"\<lbrakk>F \<in> increasing[A](r, f); mono1(A, r, B, s, g); |
|
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
80 |
refl(A, r); trans[B](s)\<rbrakk> \<Longrightarrow> F \<in> increasing[B](s, g comp f)" |
|
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
81 |
by (rule subset_increasing_comp [THEN subsetD], auto) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
82 |
|
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
83 |
lemma strict_increasing: |
| 46823 | 84 |
"increasing[nat](Le, f) \<subseteq> increasing[nat](Lt, f)" |
|
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
85 |
by (unfold increasing_def Lt_def, auto) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
86 |
|
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
87 |
lemma strict_gt_increasing: |
| 46823 | 88 |
"increasing[nat](Ge, f) \<subseteq> increasing[nat](Gt, f)" |
|
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
89 |
apply (unfold increasing_def Gt_def Ge_def, auto) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
90 |
apply (erule natE) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
91 |
apply (auto simp add: stable_def) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
92 |
done |
|
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 |
(** Increasing **) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
95 |
|
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
96 |
lemma increasing_imp_Increasing: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
97 |
"F \<in> increasing[A](r, f) \<Longrightarrow> F \<in> Increasing[A](r, f)" |
|
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
98 |
|
| 76217 | 99 |
unfolding increasing_def Increasing_def |
|
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
100 |
apply (auto intro: stable_imp_Stable) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
101 |
done |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
102 |
|
| 46823 | 103 |
lemma Increasing_type: "Increasing[A](r, f) \<subseteq> program" |
|
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
104 |
by (unfold Increasing_def, auto) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
105 |
|
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
106 |
lemma Increasing_into_program: "F \<in> Increasing[A](r, f) \<Longrightarrow> F \<in> program" |
|
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
107 |
by (unfold Increasing_def, auto) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
108 |
|
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
109 |
lemma Increasing_imp_Stable: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
110 |
"\<lbrakk>F \<in> Increasing[A](r, f); a \<in> A\<rbrakk> \<Longrightarrow> F \<in> Stable({s \<in> state. <a,f(s)>:r})"
|
|
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
111 |
by (unfold Increasing_def, blast) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
112 |
|
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
113 |
lemma IncreasingD: |
| 76214 | 114 |
"F \<in> Increasing[A](r, f) \<Longrightarrow> F \<in> program \<and> (\<exists>a. a \<in> A) \<and> (\<forall>s \<in> state. f(s):A)" |
|
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
115 |
unfolding Increasing_def |
|
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
116 |
apply (subgoal_tac "\<exists>x. x \<in> state") |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
117 |
apply (auto intro: st0_in_state) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
118 |
done |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
119 |
|
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
120 |
lemma Increasing_constant [simp]: |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
121 |
"F \<in> Increasing[A](r, \<lambda>s. c) \<longleftrightarrow> F \<in> program \<and> (c \<in> A)" |
|
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
122 |
apply (subgoal_tac "\<exists>x. x \<in> state") |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
123 |
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
|
124 |
done |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
125 |
|
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
126 |
lemma subset_Increasing_comp: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
127 |
"\<lbrakk>mono1(A, r, B, s, g); refl(A, r); trans[B](s)\<rbrakk> \<Longrightarrow> |
| 46823 | 128 |
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
|
129 |
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
|
130 |
constrains_def mono1_def metacomp_def, safe) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
131 |
apply (simp_all add: ActsD) |
| 76214 | 132 |
apply (subgoal_tac "xb \<in> state \<and> xa \<in> state") |
|
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
133 |
prefer 2 apply (simp add: ActsD) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
134 |
apply (subgoal_tac "<f (xb), f (xb) >:r") |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
135 |
prefer 2 apply (force simp add: refl_def) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
136 |
apply (rotate_tac 5) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
137 |
apply (drule_tac x = "f (xb) " in bspec) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
138 |
apply simp_all |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
139 |
apply clarify |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
140 |
apply (rotate_tac -2) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
141 |
apply (drule_tac x = act in bspec) |
| 59788 | 142 |
apply (drule_tac [2] A = "act``u" and c = xa for u in subsetD, simp_all, blast) |
143 |
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
|
144 |
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
|
145 |
apply simp_all |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
146 |
done |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
147 |
|
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
148 |
lemma imp_Increasing_comp: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
149 |
"\<lbrakk>F \<in> Increasing[A](r, f); mono1(A, r, B, s, g); refl(A, r); trans[B](s)\<rbrakk> |
|
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
150 |
\<Longrightarrow> F \<in> Increasing[B](s, g comp f)" |
|
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
151 |
apply (rule subset_Increasing_comp [THEN subsetD], auto) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
152 |
done |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
153 |
|
| 46823 | 154 |
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
|
155 |
by (unfold Increasing_def Lt_def, auto) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
156 |
|
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
157 |
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
|
158 |
apply (unfold Increasing_def Ge_def Gt_def, auto) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
159 |
apply (erule natE) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
160 |
apply (auto simp add: Stable_def) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
161 |
done |
|
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 |
(** Two-place monotone operations **) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
164 |
|
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
165 |
lemma imp_increasing_comp2: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
166 |
"\<lbrakk>F \<in> increasing[A](r, f); F \<in> increasing[B](s, g); |
|
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
167 |
mono2(A, r, B, s, C, t, h); refl(A, r); refl(B, s); trans[C](t)\<rbrakk> |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
168 |
\<Longrightarrow> F \<in> increasing[C](t, \<lambda>x. h(f(x), g(x)))" |
|
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
169 |
apply (unfold increasing_def stable_def |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
170 |
part_order_def constrains_def mono2_def, clarify, simp) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
171 |
apply clarify |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
172 |
apply (rename_tac xa xb) |
| 76214 | 173 |
apply (subgoal_tac "xb \<in> state \<and> xa \<in> state") |
|
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
174 |
prefer 2 apply (blast dest!: ActsD) |
| 76214 | 175 |
apply (subgoal_tac "<f (xb), f (xb) >:r \<and> <g (xb), g (xb) >:s") |
|
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
176 |
prefer 2 apply (force simp add: refl_def) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
177 |
apply (rotate_tac 6) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
178 |
apply (drule_tac x = "f (xb) " in bspec) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
179 |
apply (rotate_tac [2] 1) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
180 |
apply (drule_tac [2] x = "g (xb) " in bspec) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
181 |
apply simp_all |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
182 |
apply (rotate_tac -1) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
183 |
apply (drule_tac x = act in bspec) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
184 |
apply (rotate_tac [2] -3) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
185 |
apply (drule_tac [2] x = act in bspec, simp_all) |
| 59788 | 186 |
apply (drule_tac A = "act``u" and c = xa for u in subsetD) |
187 |
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
|
188 |
apply (rotate_tac -4) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
189 |
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
|
190 |
apply (rotate_tac [3] -1) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
191 |
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
|
192 |
apply simp_all |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
193 |
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
|
194 |
apply simp_all |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
195 |
done |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
196 |
|
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
197 |
lemma imp_Increasing_comp2: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
198 |
"\<lbrakk>F \<in> Increasing[A](r, f); F \<in> Increasing[B](s, g); |
|
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
199 |
mono2(A, r, B, s, C, t, h); refl(A, r); refl(B, s); trans[C](t)\<rbrakk> \<Longrightarrow> |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
200 |
F \<in> Increasing[C](t, \<lambda>x. h(f(x), g(x)))" |
|
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
201 |
apply (unfold Increasing_def stable_def |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
202 |
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
|
203 |
apply (simp_all add: ActsD) |
| 76214 | 204 |
apply (subgoal_tac "xa \<in> state \<and> x \<in> state") |
|
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
205 |
prefer 2 apply (blast dest!: ActsD) |
| 76214 | 206 |
apply (subgoal_tac "<f (xa), f (xa) >:r \<and> <g (xa), g (xa) >:s") |
|
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
207 |
prefer 2 apply (force simp add: refl_def) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
208 |
apply (rotate_tac 6) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
209 |
apply (drule_tac x = "f (xa) " in bspec) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
210 |
apply (rotate_tac [2] 1) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
211 |
apply (drule_tac [2] x = "g (xa) " in bspec) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
212 |
apply simp_all |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
213 |
apply clarify |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
214 |
apply (rotate_tac -2) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
215 |
apply (drule_tac x = act in bspec) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
216 |
apply (rotate_tac [2] -3) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
217 |
apply (drule_tac [2] x = act in bspec, simp_all) |
| 59788 | 218 |
apply (drule_tac A = "act``u" and c = x for u in subsetD) |
219 |
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
|
220 |
apply (rotate_tac -9) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
221 |
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
|
222 |
apply (rotate_tac [3] -1) |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
223 |
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
|
224 |
apply simp_all |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
225 |
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
|
226 |
apply simp_all |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
227 |
done |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14055
diff
changeset
|
228 |
|
| 14052 | 229 |
end |