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