src/HOL/Induct/Comb.thy
changeset 3309 992a25b24d0d
parent 3120 c58423c20740
child 5102 8c782c25a11e
equal deleted inserted replaced
3308:da002cef7090 3309:992a25b24d0d
    11     Reasoning with Inductively Defined Relations in the HOL Theorem Prover.
    11     Reasoning with Inductively Defined Relations in the HOL Theorem Prover.
    12     Report 265, University of Cambridge Computer Laboratory, 1992.
    12     Report 265, University of Cambridge Computer Laboratory, 1992.
    13 *)
    13 *)
    14 
    14 
    15 
    15 
    16 Comb = Trancl +
    16 Comb = Arith +
    17 
    17 
    18 (** Datatype definition of combinators S and K, with infixed application **)
    18 (** Datatype definition of combinators S and K, with infixed application **)
    19 datatype comb = K
    19 datatype comb = K
    20               | S
    20               | S
    21               | "#" comb comb (infixl 90)
    21               | "#" comb comb (infixl 90)