| author | wenzelm | 
| Mon, 11 Feb 2002 10:56:33 +0100 | |
| changeset 12873 | d7f8dfaad46d | 
| parent 11359 | 29f8b00d7e1f | 
| child 13075 | d3e1d554cd6d | 
| permissions | -rw-r--r-- | 
| 
3120
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
1  | 
(* Title: HOL/ex/Comb.thy  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
2  | 
ID: $Id$  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
3  | 
Author: Lawrence C Paulson  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
4  | 
Copyright 1996 University of Cambridge  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
5  | 
|
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
6  | 
Combinatory Logic example: the Church-Rosser Theorem  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
7  | 
Curiously, combinators do not include free variables.  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
8  | 
|
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
9  | 
Example taken from  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
10  | 
J. Camilleri and T. F. Melham.  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
11  | 
Reasoning with Inductively Defined Relations in the HOL Theorem Prover.  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
12  | 
Report 265, University of Cambridge Computer Laboratory, 1992.  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
13  | 
*)  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
14  | 
|
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
15  | 
|
| 9101 | 16  | 
Comb = Main +  | 
| 
3120
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
17  | 
|
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
18  | 
(** Datatype definition of combinators S and K, with infixed application **)  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
19  | 
datatype comb = K  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
20  | 
| S  | 
| 11359 | 21  | 
| "##" comb comb (infixl 90)  | 
| 
3120
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
22  | 
|
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
23  | 
(** Inductive definition of contractions, -1->  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
24  | 
and (multi-step) reductions, --->  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
25  | 
**)  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
26  | 
consts  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
27  | 
contract :: "(comb*comb) set"  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
28  | 
"-1->" :: [comb,comb] => bool (infixl 50)  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
29  | 
"--->" :: [comb,comb] => bool (infixl 50)  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
30  | 
|
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
31  | 
translations  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
32  | 
"x -1-> y" == "(x,y) : contract"  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
33  | 
"x ---> y" == "(x,y) : contract^*"  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
34  | 
|
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
35  | 
inductive contract  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
36  | 
intrs  | 
| 11359 | 37  | 
K "K##x##y -1-> x"  | 
38  | 
S "S##x##y##z -1-> (x##z)##(y##z)"  | 
|
39  | 
Ap1 "x-1->y ==> x##z -1-> y##z"  | 
|
40  | 
Ap2 "x-1->y ==> z##x -1-> z##y"  | 
|
| 
3120
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
41  | 
|
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
42  | 
|
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
43  | 
(** Inductive definition of parallel contractions, =1=>  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
44  | 
and (multi-step) parallel reductions, ===>  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
45  | 
**)  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
46  | 
consts  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
47  | 
parcontract :: "(comb*comb) set"  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
48  | 
"=1=>" :: [comb,comb] => bool (infixl 50)  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
49  | 
"===>" :: [comb,comb] => bool (infixl 50)  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
50  | 
|
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
51  | 
translations  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
52  | 
"x =1=> y" == "(x,y) : parcontract"  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
53  | 
"x ===> y" == "(x,y) : parcontract^*"  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
54  | 
|
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
55  | 
inductive parcontract  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
56  | 
intrs  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
57  | 
refl "x =1=> x"  | 
| 11359 | 58  | 
K "K##x##y =1=> x"  | 
59  | 
S "S##x##y##z =1=> (x##z)##(y##z)"  | 
|
60  | 
Ap "[| x=1=>y; z=1=>w |] ==> x##z =1=> y##w"  | 
|
| 
3120
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
61  | 
|
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
62  | 
|
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
63  | 
(*Misc definitions*)  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
64  | 
constdefs  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
65  | 
I :: comb  | 
| 11359 | 66  | 
"I == S##K##K"  | 
| 
3120
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
67  | 
|
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
68  | 
(*confluence; Lambda/Commutation treats this more abstractly*)  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
69  | 
  diamond   :: "('a * 'a)set => bool"	
 | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
70  | 
"diamond(r) == ALL x y. (x,y):r -->  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
71  | 
(ALL y'. (x,y'):r -->  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
72  | 
(EX z. (y,z):r & (y',z) : r))"  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
73  | 
|
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
74  | 
end  |