| author | paulson | 
| Wed, 08 May 1996 17:49:16 +0200 | |
| changeset 1733 | 89dd6ca7ee6c | 
| parent 16 | 0b033d50ca1c | 
| permissions | -rw-r--r-- | 
| 0 | 1  | 
(* Title: ZF/ex/contract.thy  | 
2  | 
ID: $Id$  | 
|
| 
16
 
0b033d50ca1c
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 
lcp 
parents: 
0 
diff
changeset
 | 
3  | 
Author: Lawrence C Paulson  | 
| 0 | 4  | 
Copyright 1993 University of Cambridge  | 
5  | 
||
6  | 
Inductive definition of (1-step) contractions and (mult-step) reductions  | 
|
7  | 
*)  | 
|
8  | 
||
9  | 
Contract0 = Comb +  | 
|
10  | 
consts  | 
|
11  | 
diamond :: "i => o"  | 
|
12  | 
I :: "i"  | 
|
13  | 
||
14  | 
contract :: "i"  | 
|
15  | 
"-1->" :: "[i,i] => o" (infixl 50)  | 
|
16  | 
"--->" :: "[i,i] => o" (infixl 50)  | 
|
17  | 
||
18  | 
parcontract :: "i"  | 
|
19  | 
"=1=>" :: "[i,i] => o" (infixl 50)  | 
|
20  | 
"===>" :: "[i,i] => o" (infixl 50)  | 
|
21  | 
||
22  | 
translations  | 
|
23  | 
"p -1-> q" == "<p,q> : contract"  | 
|
24  | 
"p ---> q" == "<p,q> : contract^*"  | 
|
25  | 
"p =1=> q" == "<p,q> : parcontract"  | 
|
26  | 
"p ===> q" == "<p,q> : parcontract^+"  | 
|
27  | 
||
28  | 
rules  | 
|
29  | 
||
30  | 
diamond_def "diamond(r) == ALL x y. <x,y>:r --> \  | 
|
31  | 
\ (ALL y'. <x,y'>:r --> \  | 
|
32  | 
\ (EX z. <y,z>:r & <y',z> : r))"  | 
|
33  | 
||
34  | 
I_def "I == S#K#K"  | 
|
35  | 
||
36  | 
end  |