| author | berghofe | 
| Mon, 17 Dec 2007 18:24:44 +0100 | |
| changeset 25674 | b04508c59b9d | 
| parent 25135 | 4f8176c940cf | 
| child 27185 | 0407630909ef | 
| permissions | -rw-r--r-- | 
| 17291 | 1  | 
(* $Id$ *)  | 
2  | 
||
| 2570 | 3  | 
(* Specification of the following loop back device  | 
4  | 
||
5  | 
||
| 17291 | 6  | 
g  | 
| 2570 | 7  | 
--------------------  | 
8  | 
| ------- |  | 
|
9  | 
x | | | | y  | 
|
| 17291 | 10  | 
------|---->| |------| ----->  | 
| 2570 | 11  | 
| z | f | z |  | 
12  | 
| -->| |--- |  | 
|
13  | 
| | | | | |  | 
|
14  | 
| | ------- | |  | 
|
15  | 
| | | |  | 
|
16  | 
| <-------------- |  | 
|
| 17291 | 17  | 
| |  | 
| 2570 | 18  | 
--------------------  | 
19  | 
||
20  | 
||
21  | 
First step: Notation in Agent Network Description Language (ANDL)  | 
|
22  | 
-----------------------------------------------------------------  | 
|
23  | 
||
24  | 
agent f  | 
|
| 17291 | 25  | 
        input  channel i1:'b i2: ('b,'c) tc
 | 
26  | 
        output channel o1:'c o2: ('b,'c) tc
 | 
|
| 2570 | 27  | 
is  | 
| 17291 | 28  | 
Rf(i1,i2,o1,o2) (left open in the example)  | 
| 2570 | 29  | 
end f  | 
30  | 
||
31  | 
agent g  | 
|
| 17291 | 32  | 
input channel x:'b  | 
33  | 
output channel y:'c  | 
|
| 2570 | 34  | 
is network  | 
| 17291 | 35  | 
<y,z> = f$<x,z>  | 
| 2570 | 36  | 
end network  | 
37  | 
end g  | 
|
38  | 
||
39  | 
||
40  | 
Remark: the type of the feedback depends at most on the types of the input and  | 
|
41  | 
output of g. (No type miracles inside g)  | 
|
42  | 
||
43  | 
Second step: Translation of ANDL specification to HOLCF Specification  | 
|
44  | 
---------------------------------------------------------------------  | 
|
45  | 
||
46  | 
Specification of agent f ist translated to predicate is_f  | 
|
47  | 
||
| 17291 | 48  | 
is_f :: ('b stream * ('b,'c) tc stream ->
 | 
49  | 
                'c stream * ('b,'c) tc stream) => bool
 | 
|
| 2570 | 50  | 
|
| 17291 | 51  | 
is_f f = !i1 i2 o1 o2.  | 
52  | 
f$<i1,i2> = <o1,o2> --> Rf(i1,i2,o1,o2)  | 
|
| 2570 | 53  | 
|
54  | 
Specification of agent g is translated to predicate is_g which uses  | 
|
55  | 
predicate is_net_g  | 
|
56  | 
||
57  | 
is_net_g :: ('b stream * ('b,'c) tc stream -> 'c stream * ('b,'c) tc stream) =>
 | 
|
| 17291 | 58  | 
'b stream => 'c stream => bool  | 
| 2570 | 59  | 
|
| 17291 | 60  | 
is_net_g f x y =  | 
61  | 
? z. <y,z> = f$<x,z> &  | 
|
62  | 
!oy hz. <oy,hz> = f$<x,hz> --> z << hz  | 
|
| 2570 | 63  | 
|
64  | 
||
65  | 
is_g :: ('b stream -> 'c stream) => bool
 | 
|
66  | 
||
| 10835 | 67  | 
is_g g = ? f. is_f f & (!x y. g$x = y --> is_net_g f x y  | 
| 17291 | 68  | 
|
| 2570 | 69  | 
Third step: (show conservativity)  | 
70  | 
-----------  | 
|
71  | 
||
72  | 
Suppose we have a model for the theory TH1 which contains the axiom  | 
|
73  | 
||
| 17291 | 74  | 
? f. is_f f  | 
| 2570 | 75  | 
|
76  | 
In this case there is also a model for the theory TH2 that enriches TH1 by  | 
|
77  | 
axiom  | 
|
78  | 
||
| 17291 | 79  | 
? g. is_g g  | 
| 2570 | 80  | 
|
81  | 
The result is proved by showing that there is a definitional extension  | 
|
82  | 
that extends TH1 by a definition of g.  | 
|
83  | 
||
84  | 
||
85  | 
We define:  | 
|
86  | 
||
| 17291 | 87  | 
def_g g =  | 
88  | 
(? f. is_f f &  | 
|
89  | 
g = (LAM x. cfst$(f$<x,fix$(LAM k.csnd$(f$<x,k>))>)) )  | 
|
90  | 
||
| 2570 | 91  | 
Now we prove:  | 
92  | 
||
| 17291 | 93  | 
(? f. is_f f ) --> (? g. is_g g)  | 
| 2570 | 94  | 
|
95  | 
using the theorems  | 
|
96  | 
||
| 17291 | 97  | 
loopback_eq) def_g = is_g (real work)  | 
| 2570 | 98  | 
|
| 17291 | 99  | 
L1) (? f. is_f f ) --> (? g. def_g g) (trivial)  | 
| 2570 | 100  | 
|
101  | 
*)  | 
|
102  | 
||
| 17291 | 103  | 
theory Focus_ex  | 
104  | 
imports Stream  | 
|
105  | 
begin  | 
|
| 2570 | 106  | 
|
| 17291 | 107  | 
typedecl ('a, 'b) tc
 | 
108  | 
arities tc:: (pcpo, pcpo) pcpo  | 
|
| 2570 | 109  | 
|
| 
25135
 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 
wenzelm 
parents: 
21404 
diff
changeset
 | 
110  | 
axiomatization  | 
| 19763 | 111  | 
  Rf :: "('b stream * ('b,'c) tc stream * 'c stream * ('b,'c) tc stream) => bool"
 | 
| 2570 | 112  | 
|
| 19763 | 113  | 
definition  | 
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19763 
diff
changeset
 | 
114  | 
  is_f :: "('b stream * ('b,'c) tc stream -> 'c stream * ('b,'c) tc stream) => bool" where
 | 
| 19763 | 115  | 
"is_f f = (!i1 i2 o1 o2. f$<i1,i2> = <o1,o2> --> Rf(i1,i2,o1,o2))"  | 
| 2570 | 116  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19763 
diff
changeset
 | 
117  | 
definition  | 
| 19763 | 118  | 
  is_net_g :: "('b stream *('b,'c) tc stream -> 'c stream * ('b,'c) tc stream) =>
 | 
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19763 
diff
changeset
 | 
119  | 
'b stream => 'c stream => bool" where  | 
| 19763 | 120  | 
"is_net_g f x y == (? z.  | 
| 17291 | 121  | 
<y,z> = f$<x,z> &  | 
122  | 
(!oy hz. <oy,hz> = f$<x,hz> --> z << hz))"  | 
|
| 2570 | 123  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19763 
diff
changeset
 | 
124  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19763 
diff
changeset
 | 
125  | 
  is_g :: "('b stream -> 'c stream) => bool" where
 | 
| 19763 | 126  | 
"is_g g == (? f. is_f f & (!x y. g$x = y --> is_net_g f x y))"  | 
| 2570 | 127  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19763 
diff
changeset
 | 
128  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19763 
diff
changeset
 | 
129  | 
  def_g :: "('b stream -> 'c stream) => bool" where
 | 
| 19763 | 130  | 
"def_g g == (? f. is_f f & g = (LAM x. cfst$(f$<x,fix$(LAM k. csnd$(f$<x,k>))>)))"  | 
| 17291 | 131  | 
|
| 19742 | 132  | 
|
133  | 
(* first some logical trading *)  | 
|
134  | 
||
135  | 
lemma lemma1:  | 
|
136  | 
"is_g(g) =  | 
|
137  | 
(? f. is_f(f) & (!x.(? z. <g$x,z> = f$<x,z> &  | 
|
138  | 
(! w y. <y,w> = f$<x,w> --> z << w))))"  | 
|
| 19763 | 139  | 
apply (simp add: is_g_def is_net_g_def)  | 
| 19742 | 140  | 
apply fast  | 
141  | 
done  | 
|
142  | 
||
143  | 
lemma lemma2:  | 
|
144  | 
"(? f. is_f(f) & (!x. (? z. <g$x,z> = f$<x,z> &  | 
|
145  | 
(!w y. <y,w> = f$<x,w> --> z << w))))  | 
|
146  | 
=  | 
|
147  | 
(? f. is_f(f) & (!x. ? z.  | 
|
148  | 
g$x = cfst$(f$<x,z>) &  | 
|
149  | 
z = csnd$(f$<x,z>) &  | 
|
150  | 
(! w y. <y,w> = f$<x,w> --> z << w)))"  | 
|
151  | 
apply (rule iffI)  | 
|
152  | 
apply (erule exE)  | 
|
153  | 
apply (rule_tac x = "f" in exI)  | 
|
154  | 
apply (erule conjE)+  | 
|
155  | 
apply (erule conjI)  | 
|
156  | 
apply (intro strip)  | 
|
157  | 
apply (erule allE)  | 
|
158  | 
apply (erule exE)  | 
|
159  | 
apply (rule_tac x = "z" in exI)  | 
|
160  | 
apply (erule conjE)+  | 
|
161  | 
apply (rule conjI)  | 
|
162  | 
apply (rule_tac [2] conjI)  | 
|
163  | 
prefer 3 apply (assumption)  | 
|
164  | 
apply (drule sym)  | 
|
165  | 
apply (tactic "asm_simp_tac HOLCF_ss 1")  | 
|
166  | 
apply (drule sym)  | 
|
167  | 
apply (tactic "asm_simp_tac HOLCF_ss 1")  | 
|
168  | 
apply (erule exE)  | 
|
169  | 
apply (rule_tac x = "f" in exI)  | 
|
170  | 
apply (erule conjE)+  | 
|
171  | 
apply (erule conjI)  | 
|
172  | 
apply (intro strip)  | 
|
173  | 
apply (erule allE)  | 
|
174  | 
apply (erule exE)  | 
|
175  | 
apply (rule_tac x = "z" in exI)  | 
|
176  | 
apply (erule conjE)+  | 
|
177  | 
apply (rule conjI)  | 
|
178  | 
prefer 2 apply (assumption)  | 
|
179  | 
apply (rule trans)  | 
|
180  | 
apply (rule_tac [2] surjective_pairing_Cprod2)  | 
|
181  | 
apply (erule subst)  | 
|
182  | 
apply (erule subst)  | 
|
183  | 
apply (rule refl)  | 
|
184  | 
done  | 
|
185  | 
||
186  | 
lemma lemma3: "def_g(g) --> is_g(g)"  | 
|
| 19763 | 187  | 
apply (tactic {* simp_tac (HOL_ss addsimps [thm "def_g_def", thm "lemma1", thm "lemma2"]) 1 *})
 | 
| 19742 | 188  | 
apply (rule impI)  | 
189  | 
apply (erule exE)  | 
|
190  | 
apply (rule_tac x = "f" in exI)  | 
|
191  | 
apply (erule conjE)+  | 
|
192  | 
apply (erule conjI)  | 
|
193  | 
apply (intro strip)  | 
|
194  | 
apply (rule_tac x = "fix$ (LAM k. csnd$ (f$<x,k>))" in exI)  | 
|
195  | 
apply (rule conjI)  | 
|
196  | 
apply (tactic "asm_simp_tac HOLCF_ss 1")  | 
|
197  | 
apply (rule trans)  | 
|
198  | 
apply (rule_tac [2] surjective_pairing_Cprod2)  | 
|
199  | 
apply (rule cfun_arg_cong)  | 
|
200  | 
apply (rule trans)  | 
|
201  | 
apply (rule fix_eq)  | 
|
202  | 
apply (simp (no_asm))  | 
|
203  | 
apply (intro strip)  | 
|
204  | 
apply (rule fix_least)  | 
|
205  | 
apply (simp (no_asm))  | 
|
206  | 
apply (erule exE)  | 
|
207  | 
apply (drule sym)  | 
|
208  | 
back  | 
|
209  | 
apply simp  | 
|
210  | 
done  | 
|
211  | 
||
212  | 
lemma lemma4: "is_g(g) --> def_g(g)"  | 
|
213  | 
apply (tactic {* simp_tac (HOL_ss delsimps (thms "ex_simps" @ thms "all_simps")
 | 
|
| 19763 | 214  | 
addsimps [thm "lemma1", thm "lemma2", thm "def_g_def"]) 1 *})  | 
| 19742 | 215  | 
apply (rule impI)  | 
216  | 
apply (erule exE)  | 
|
217  | 
apply (rule_tac x = "f" in exI)  | 
|
218  | 
apply (erule conjE)+  | 
|
219  | 
apply (erule conjI)  | 
|
220  | 
apply (rule ext_cfun)  | 
|
221  | 
apply (erule_tac x = "x" in allE)  | 
|
222  | 
apply (erule exE)  | 
|
223  | 
apply (erule conjE)+  | 
|
224  | 
apply (subgoal_tac "fix$ (LAM k. csnd$ (f$<x, k>)) = z")  | 
|
225  | 
apply simp  | 
|
226  | 
apply (subgoal_tac "! w y. f$<x, w> = <y, w> --> z << w")  | 
|
227  | 
apply (rule sym)  | 
|
228  | 
apply (rule fix_eqI)  | 
|
229  | 
apply simp  | 
|
230  | 
apply (rule allI)  | 
|
231  | 
apply (tactic "simp_tac HOLCF_ss 1")  | 
|
232  | 
apply (intro strip)  | 
|
233  | 
apply (subgoal_tac "f$<x, za> = <cfst$ (f$<x,za>) ,za>")  | 
|
234  | 
apply fast  | 
|
235  | 
apply (rule trans)  | 
|
236  | 
apply (rule surjective_pairing_Cprod2 [symmetric])  | 
|
237  | 
apply (erule cfun_arg_cong)  | 
|
238  | 
apply (intro strip)  | 
|
239  | 
apply (erule allE)+  | 
|
240  | 
apply (erule mp)  | 
|
241  | 
apply (erule sym)  | 
|
242  | 
done  | 
|
243  | 
||
244  | 
(* now we assemble the result *)  | 
|
245  | 
||
246  | 
lemma loopback_eq: "def_g = is_g"  | 
|
247  | 
apply (rule ext)  | 
|
248  | 
apply (rule iffI)  | 
|
249  | 
apply (erule lemma3 [THEN mp])  | 
|
250  | 
apply (erule lemma4 [THEN mp])  | 
|
251  | 
done  | 
|
252  | 
||
253  | 
lemma L2:  | 
|
254  | 
"(? f.  | 
|
255  | 
  is_f(f::'b stream * ('b,'c) tc stream -> 'c stream * ('b,'c) tc stream))
 | 
|
256  | 
-->  | 
|
257  | 
(? g. def_g(g::'b stream -> 'c stream ))"  | 
|
| 19763 | 258  | 
apply (simp add: def_g_def)  | 
| 19742 | 259  | 
done  | 
260  | 
||
261  | 
theorem conservative_loopback:  | 
|
262  | 
"(? f.  | 
|
263  | 
  is_f(f::'b stream * ('b,'c) tc stream -> 'c stream * ('b,'c) tc stream))
 | 
|
264  | 
-->  | 
|
265  | 
(? g. is_g(g::'b stream -> 'c stream ))"  | 
|
266  | 
apply (rule loopback_eq [THEN subst])  | 
|
267  | 
apply (rule L2)  | 
|
268  | 
done  | 
|
| 2570 | 269  | 
|
270  | 
end  |