src/HOLCF/ex/Focus_ex.thy
author nipkow
Tue, 09 Jan 2001 15:36:30 +0100
changeset 10835 f4745d77e620
parent 3842 b55686a7b22c
child 17291 94f6113fe9ed
permissions -rw-r--r--
` -> $
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     1
(* Specification of the following loop back device
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     2
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     3
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     4
          g 
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     5
           --------------------
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     6
          |      -------       |
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     7
       x  |     |       |      |  y
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     8
    ------|---->|       |------| ----->        
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     9
          |  z  |   f   | z    |
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    10
          |  -->|       |---   |
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    11
          | |   |       |   |  |
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    12
          | |    -------    |  |
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    13
          | |               |  |
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    14
          |  <--------------   |
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    15
          |                    | 
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    16
           --------------------
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    17
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    18
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    19
First step: Notation in Agent Network Description Language (ANDL)
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    20
-----------------------------------------------------------------
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    21
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    22
agent f
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    23
	input  channel i1:'b i2: ('b,'c) tc
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    24
	output channel o1:'c o2: ('b,'c) tc
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    25
is
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    26
	Rf(i1,i2,o1,o2)  (left open in the example)
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    27
end f
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    28
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    29
agent g
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    30
	input  channel x:'b 
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    31
	output channel y:'c 
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    32
is network
10835
nipkow
parents: 3842
diff changeset
    33
	<y,z> = f$<x,z>
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    34
end network
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    35
end g
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    36
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    37
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    38
Remark: the type of the feedback depends at most on the types of the input and
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    39
        output of g. (No type miracles inside g)
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    40
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    41
Second step: Translation of ANDL specification to HOLCF Specification
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    42
---------------------------------------------------------------------
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    43
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    44
Specification of agent f ist translated to predicate is_f
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    45
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    46
is_f :: ('b stream * ('b,'c) tc stream -> 
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    47
		'c stream * ('b,'c) tc stream) => bool
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    48
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    49
is_f f  = !i1 i2 o1 o2. 
10835
nipkow
parents: 3842
diff changeset
    50
	f$<i1,i2> = <o1,o2> --> Rf(i1,i2,o1,o2)
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    51
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    52
Specification of agent g is translated to predicate is_g which uses
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    53
predicate is_net_g
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    54
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    55
is_net_g :: ('b stream * ('b,'c) tc stream -> 'c stream * ('b,'c) tc stream) =>
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    56
	    'b stream => 'c stream => bool
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    57
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    58
is_net_g f x y = 
10835
nipkow
parents: 3842
diff changeset
    59
	? z. <y,z> = f$<x,z> & 
nipkow
parents: 3842
diff changeset
    60
	!oy hz. <oy,hz> = f$<x,hz> --> z << hz 
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    61
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    62
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    63
is_g :: ('b stream -> 'c stream) => bool
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    64
10835
nipkow
parents: 3842
diff changeset
    65
is_g g  = ? f. is_f f  & (!x y. g$x = y --> is_net_g f x y
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    66
	  
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    67
Third step: (show conservativity)
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    68
-----------
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    69
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    70
Suppose we have a model for the theory TH1 which contains the axiom
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    71
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    72
	? f. is_f f 
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    73
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    74
In this case there is also a model for the theory TH2 that enriches TH1 by
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    75
axiom
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    76
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    77
	? g. is_g g 
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    78
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    79
The result is proved by showing that there is a definitional extension
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    80
that extends TH1 by a definition of g.
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    81
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    82
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    83
We define:
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    84
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    85
def_g g  = 
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    86
         (? f. is_f f  & 
10835
nipkow
parents: 3842
diff changeset
    87
	      g = (LAM x. cfst$(f$<x,fix$(LAM k.csnd$(f$<x,k>))>)) )
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    88
	
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    89
Now we prove:
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    90
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    91
	(? f. is_f f ) --> (? g. is_g g) 
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    92
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    93
using the theorems
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    94
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    95
loopback_eq)	def_g = is_g			(real work) 
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    96
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    97
L1)		(? f. is_f f ) --> (? g. def_g g)  (trivial)
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    98
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    99
*)
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   100
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   101
Focus_ex = Stream +
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   102
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   103
types  tc 2
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   104
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   105
arities tc:: (pcpo,pcpo)pcpo
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   106
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   107
consts
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   108
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   109
is_f     ::
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   110
 "('b stream * ('b,'c) tc stream -> 'c stream * ('b,'c) tc stream) => bool"
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   111
is_net_g :: "('b stream *('b,'c) tc stream -> 'c stream * ('b,'c) tc stream) =>
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   112
	    'b stream => 'c stream => bool"
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   113
is_g     :: "('b stream -> 'c stream) => bool"
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   114
def_g    :: "('b stream -> 'c stream) => bool"
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   115
Rf	 :: 
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   116
"('b stream * ('b,'c) tc stream * 'c stream * ('b,'c) tc stream) => bool"
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   117
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   118
defs
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   119
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   120
is_f		"is_f f == (!i1 i2 o1 o2.
10835
nipkow
parents: 3842
diff changeset
   121
			f$<i1,i2> = <o1,o2> --> Rf(i1,i2,o1,o2))"
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   122
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   123
is_net_g	"is_net_g f x y == (? z. 
10835
nipkow
parents: 3842
diff changeset
   124
			<y,z> = f$<x,z> &
nipkow
parents: 3842
diff changeset
   125
			(!oy hz. <oy,hz> = f$<x,hz> --> z << hz))" 
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   126
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   127
is_g		"is_g g  == (? f.
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   128
			is_f f  & 
10835
nipkow
parents: 3842
diff changeset
   129
			(!x y. g$x = y --> is_net_g f x y))"
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   130
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   131
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   132
def_g		"def_g g == (? f.
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   133
			is_f f  & 
10835
nipkow
parents: 3842
diff changeset
   134
	      		g = (LAM x. cfst$(f$<x,fix$(LAM  k. csnd$(f$<x,k>))>)))" 
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   135
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   136
end