| author | wenzelm | 
| Fri, 28 Sep 2001 19:23:58 +0200 | |
| changeset 11633 | c8945e0dc00b | 
| parent 11354 | 9b80fe19407f | 
| child 12867 | 5c900a821a7c | 
| permissions | -rw-r--r-- | 
| 1478 | 1 | (* Title: ZF/ex/ramsey.thy | 
| 0 | 2 | ID: $Id$ | 
| 1478 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 0 | 4 | Copyright 1992 University of Cambridge | 
| 5 | ||
| 6 | Ramsey's Theorem (finite exponent 2 version) | |
| 7 | ||
| 8 | Based upon the article | |
| 9 | D Basin and M Kaufmann, | |
| 10 | The Boyer-Moore Prover and Nuprl: An Experimental Comparison. | |
| 11 | In G Huet and G Plotkin, editors, Logical Frameworks. | |
| 12 | (CUP, 1991), pages 89--119 | |
| 13 | ||
| 14 | See also | |
| 15 | M Kaufmann, | |
| 16 | An example in NQTHM: Ramsey's Theorem | |
| 17 | Internal Note, Computational Logic, Inc., Austin, Texas 78703 | |
| 18 | Available from the author: kaufmann@cli.com | |
| 19 | *) | |
| 20 | ||
| 11354 
9b80fe19407f
examples files start from Main instead of various ZF theories
 paulson parents: 
11316diff
changeset | 21 | Ramsey = Main + | 
| 0 | 22 | consts | 
| 1478 | 23 | Symmetric :: i=>o | 
| 24 | Atleast :: [i,i]=>o | |
| 25 | Clique,Indept,Ramsey :: [i,i,i]=>o | |
| 0 | 26 | |
| 753 | 27 | defs | 
| 0 | 28 | |
| 29 | Symmetric_def | |
| 11316 | 30 | "Symmetric(E) == (\\<forall>x y. <x,y>:E --> <y,x>:E)" | 
| 0 | 31 | |
| 32 | Clique_def | |
| 11316 | 33 | "Clique(C,V,E) == (C \\<subseteq> V) & (\\<forall>x \\<in> C. \\<forall>y \\<in> C. x\\<noteq>y --> <x,y> \\<in> E)" | 
| 0 | 34 | |
| 35 | Indept_def | |
| 11316 | 36 | "Indept(I,V,E) == (I \\<subseteq> V) & (\\<forall>x \\<in> I. \\<forall>y \\<in> I. x\\<noteq>y --> <x,y> \\<notin> E)" | 
| 0 | 37 | |
| 38 | Atleast_def | |
| 11316 | 39 | "Atleast(n,S) == (\\<exists>f. f \\<in> inj(n,S))" | 
| 0 | 40 | |
| 41 | Ramsey_def | |
| 11316 | 42 | "Ramsey(n,i,j) == \\<forall>V E. Symmetric(E) & Atleast(n,V) --> | 
| 43 | (\\<exists>C. Clique(C,V,E) & Atleast(i,C)) | | |
| 44 | (\\<exists>I. Indept(I,V,E) & Atleast(j,I))" | |
| 0 | 45 | |
| 46 | end |