src/ZF/ex/Comb.thy
author lcp
Tue, 29 Nov 1994 00:31:31 +0100
changeset 753 ec86863e87c8
parent 515 abcc438e7c27
child 810 91c68f74f458
permissions -rw-r--r--
replaced "rules" by "defs"
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     1
(*  Title: 	ZF/ex/Comb.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     2
    ID:         $Id$
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     3
    Author: 	Lawrence C Paulson
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     4
    Copyright   1994  University of Cambridge
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     5
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     6
Combinatory Logic example: the Church-Rosser Theorem
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     7
Curiously, combinators do not include free variables.
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     8
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     9
Example taken from
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    10
    J. Camilleri and T. F. Melham.
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    11
    Reasoning with Inductively Defined Relations in the HOL Theorem Prover.
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    12
    Report 265, University of Cambridge Computer Laboratory, 1992.
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    13
*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    14
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    15
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    16
Comb = Univ + "Datatype" +
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    17
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    18
(** Datatype definition of combinators S and K, with infixed application **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    19
consts comb :: "i"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    20
datatype (* <= "univ(0)" *)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    21
  "comb" = K
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    22
         | S
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    23
         | "#" ("p: comb", "q: comb")   (infixl 90)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    24
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    25
(** Inductive definition of contractions, -1->
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    26
             and (multi-step) reductions, --->
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    27
**)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    28
consts
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    29
  contract  :: "i"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    30
  "-1->"    :: "[i,i] => o"    			(infixl 50)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    31
  "--->"    :: "[i,i] => o"    			(infixl 50)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    32
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    33
translations
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    34
  "p -1-> q" == "<p,q> : contract"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    35
  "p ---> q" == "<p,q> : contract^*"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    36
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    37
inductive
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    38
  domains   "contract" <= "comb*comb"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    39
  intrs
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    40
    K     "[| p:comb;  q:comb |] ==> K#p#q -1-> p"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    41
    S     "[| p:comb;  q:comb;  r:comb |] ==> S#p#q#r -1-> (p#r)#(q#r)"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    42
    Ap1   "[| p-1->q;  r:comb |] ==> p#r -1-> q#r"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    43
    Ap2   "[| p-1->q;  r:comb |] ==> r#p -1-> r#q"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    44
  type_intrs "comb.intrs"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    45
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    46
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    47
(** Inductive definition of parallel contractions, =1=>
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    48
             and (multi-step) parallel reductions, ===>
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    49
**)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    50
consts
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    51
  parcontract :: "i"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    52
  "=1=>"    :: "[i,i] => o"    			(infixl 50)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    53
  "===>"    :: "[i,i] => o"    			(infixl 50)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    54
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    55
translations
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    56
  "p =1=> q" == "<p,q> : parcontract"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    57
  "p ===> q" == "<p,q> : parcontract^+"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    58
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    59
inductive
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    60
  domains   "parcontract" <= "comb*comb"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    61
  intrs
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    62
    refl  "[| p:comb |] ==> p =1=> p"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    63
    K     "[| p:comb;  q:comb |] ==> K#p#q =1=> p"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    64
    S     "[| p:comb;  q:comb;  r:comb |] ==> S#p#q#r =1=> (p#r)#(q#r)"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    65
    Ap    "[| p=1=>q;  r=1=>s |] ==> p#r =1=> q#s"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    66
  type_intrs "comb.intrs"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    67
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    68
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    69
(*Misc definitions*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    70
consts
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    71
  diamond   :: "i => o"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    72
  I         :: "i"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    73
753
ec86863e87c8 replaced "rules" by "defs"
lcp
parents: 515
diff changeset
    74
defs
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    75
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    76
  diamond_def "diamond(r) == ALL x y. <x,y>:r --> \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    77
\                            (ALL y'. <x,y'>:r --> \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    78
\                                 (EX z. <y,z>:r & <y',z> : r))"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    79
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    80
  I_def       "I == S#K#K"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    81
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    82
end