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