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