I have modified the grammar for idts (sequences of identifiers with optional
type annotations). idts are generally used as in abstractions, be it
lambda-abstraction or quantifiers. It now has roughly the form
idts = pttrn*
pttrn = idt
where pttrn is a new nonterminal (type) not used anywhere else.
This means that the Pure syntax for idts is in fact unchanged.
The point is that the new nontermianl pttrn allows later extensions of this
syntax. (See, for example, HOL/Prod.thy).
The name idts is not quite accurate any longer and may become downright
confusing once pttrn has been extended. Something should be done about this,
in particular wrt to the manual.
(* Title: 91/Modal/modal0
ID: $Id$
Author: Martin Coen
Copyright 1991 University of Cambridge
*)
Modal0 = LK +
consts
box :: "o=>o" ("[]_" [50] 50)
dia :: "o=>o" ("<>_" [50] 50)
"--<",">-<" :: "[o,o]=>o" (infixr 25)
"@Lstar" :: "[sequence,sequence]=>prop" ("(_)|L>(_)" [6,6] 5)
"@Rstar" :: "[sequence,sequence]=>prop" ("(_)|R>(_)" [6,6] 5)
Lstar,Rstar :: "[sobj=>sobj,sobj=>sobj]=>prop"
rules
(* Definitions *)
strimp_def "P --< Q == [](P --> Q)"
streqv_def "P >-< Q == (P --< Q) & (Q --< P)"
end
ML
local
val Lstar = "Lstar";
val Rstar = "Rstar";
val SLstar = "@Lstar";
val SRstar = "@Rstar";
fun star_tr c [s1,s2] = Const(c,dummyT)$LK.seq_tr1 s1$LK.seq_tr1 s2;
fun star_tr' c [Abs(_,_,s1),Abs(_,_,s2)] =
Const(c,dummyT) $ LK.seq_tr1' s1 $ LK.seq_tr1' s2;
in
val parse_translation = [(SLstar,star_tr Lstar), (SRstar,star_tr Rstar)];
val print_translation = [(Lstar,star_tr' SLstar), (Rstar,star_tr' SRstar)]
end;