! <command>
#define <pattern> := <term> .
**
, **>
--
, -->
.
=
=(n)=>
, =(n,m)=>
, =()=>
=*=
=/=
==
==>
? [<term>]
?apropos <term> [<term> ...]
?com [ <term> ]
[
accept =*= proof
switchall axioms
switchalways memo
switch:apply (<tactic> ...) [to <goal-name>]
apply <action> [ <subst> ] <range> <selection>
:auto
auto context
switchautoload <module-name> <file-name>
ax [ <label-exp> ] <term> = <term>
.axioms { <decls> }
:backward equation|rule
bax [ <label-exp> ] <term> = <term>
.bceq [ <label-exp> ] <term> = <term> if <boolterm> .
bcrule [ <label-exp> ] <term> => <term> if <term> .
bctrans [ <label-exp> ] <term> => <term> if <bool> .
beq [ <label-exp> ] <term> = <term> .
bgoal <term> .
{bguess | :bguess} {imply|and|or} [ with <predicate name> ]
binspect [in <module-name> :] <boolean-term> .
:binspect [in <goal-name> :] <boolean-term> .
bop <op-spec> : <sorts> -> <sort>
bpred <op-spec> : <sorts>
breduce [ in <mod-exp> : ] <term> .
{bresolve | :bresolve} [<limit>] [all]
brule [ <label-exp> ] <term> => <term> .
{bshow | :bshow} [{ tree | grind }]
bsort token-predicate creater printer term-predicate
btrans [ <label-exp> ] <term> => <term> .
cbred [ in <mod-exp> :] <term> .
cd <dirname>
ceq [ <label-exp> ] <term> = <term> if <boolterm> .
check <options>
check <something>
switchchoose <selection>
clause <term> .
clean memo
clean memo
switchclose
commands
cond limit
switchcont
:cp { "[" <label> "]" | "(" <sentence> . ")" } >< { "[" <label> "]" | "(" <sentence> .")" }
crule [ <label-exp> ] <term> => <term> if <term> .
:csp { eq [ <label-exp>] <term> = <term> . ...}
:csp- { eq [ <label-exp>] <term> = <term> . ...}
:ctf { eq [ <label-exp> ] <term> = <term> .}
:ctf- { eq [ <label-exp> ] <term> = <term> .}
ctrans [ <label-exp> ] <term> => <term> if <term> .
db reset
:def <symbol> = { <ctf> | <csp> | <init> }
demod
:describe proof
describe <something>
dirs
dribble { <file-name> | .}
:embed (<label> ... <label>) as <module_name>
eof
eq [ <label-exp> ] <term> = <term> .
:equation
esc return
exec limit
switchexec trace
switchexec! [ in <mod-exp> : ] <term> .
execute [ in <mod-exp> : ] <term> .
extending ( <modexp> )
find {+rule | -rule}
find all rules
switchflag(<name>, { on | off })
full reset
gendoc <pathname>
:goal { <sentence> . ... }
goal <term> .
:imp "[" <label> "]" by "{" <variable> <- <term>; ..."}"
imports { <import-decl> }
include BOOL
switchinclude RWL
switchincluding ( <modexp> )
:ind { on (<variable> ...) | '{' on (<variable> ...) base (<Term> . ... <Term> .) step (<Term> . ... <Term> .) '}'
init [as <name>] { "[" <label> "]" | "(" <sentence> "")} by "{" <variable> <- <term>; ... "}"
:init [as <name>] { "[" <label> "]" | "(" <sentence> "")} by "{" <variable> <- <term>; ... "}"
input <pathname>
inspect <term>
:is
let <identifier> = <term> .
lex (<op>, ..., <op>)
libpath
switchlisp
lispq
list { axiom | sos | usable | flag | param | option | demod }
look up <something>
ls <pathname>
make <mod_name> ( <mod_exp> )
match <term_spec> to <pattern> .
memo
switch[sys:]module[!|*] <modname> [ ( <params> ) ] [ <principal_sort_spec> ] { mod_elements ... }
module expression
names <mod-exp>
.no autoload <module-name>
:normalize { on | off}
op <op-spec> : <sorts> -> <sort> { <attribute-list> }
open <mod_exp> .
operator attributes
operator precedence
option { reset | = <name> }
:order (<op>, ..., <op>)
param(<name>, <value>)
parameterized module
parse [ in <mod-exp> : ] <term> .
parse normalize
switchpopd
pred <op-spec> : <sorts>
prelude <file>
print depth
switchprint mode
switchprint trs
switchprotect <module-name>
protecting ( <modexp> )
provide <feature>
pushd <directory>
pvar <var-name> : <sort-name>
pwd
qualified term
quiet
switchquit
{ :red | :exec | :bred } [in <goal-name> :] <term> .
reduce [ in <mod-exp> : ] <term> .
reduce conditions
switchregularize <mod-name>
regularize signature
switchrequire <feature> [ <pathname> ]
:reset
reset
resolve {. | <file-path> }
restore <pathname>
rewrite limit
switch:roll back
:rule
rule [ <label-exp> ] <term> => <term> .
save <pathname>
save-option <name>
scase (<term>) in (<mod-exp>) as <name> { <decl> ..} : <term> .
search predicates
:select <goal-name>
select <mod_exp> .
:set(<name>, { on | off | show })
set <name> [option] <value>
:show goal|unproved|proof|discharged
show <something>
show mode
switchsigmatch (<mod-exp>) to (<mod-exp>)
signature { <sig-decl> }
sos { = | + | - } { <clause> , ... }
:spoiler { on | off}
start <term> .
statistics
switchstep
switchstop
stop pattern
switch:theory <op_name> : <arity> -> <coarity> { assoc | comm | id: <term> }
trace [whole]
switchtrans [ <label-exp> ] <term> => <term> .
unprotect <module-name>
:use (<label> ... <label>)
using ( <modexp> )
var <var-name> : <sort-name>
:verbose { on | off }
verbose
switchversion
view <name> from <modname> to <modname> { <viewelems> }
This manual introduces the language CafeOBJ. Is is a reference manual with the aim to document the current status of the language, and not targeting at an exhaustive presentation of the mathematical and logical background. Still, the next section will give a short summary of the underlying formal approach and carry references for those in search for details.
The manual is structured into three parts. The first one being this introduction, the second one being the presentation of basic concepts of CafeOBJ by providing a simple protocol which will get specified and verified. Although the second part tries to give a view onto the core features and their usage, it should not be considered a course in CafeOBJ, and cannot replace a proper introduction to the language. The CafeOBJ distribution also includes a user manual. This user manual is slightly outdated with respect to the current status of the language, but is targeting those without and prior knowledge of CafeOBJ.
Finally, the last part consists of explanations of all current language elements in alphabetic order. This includes several higher level concepts, as well as heavy cross-referencing.
While we hope that this manual and the introductory part helps beginners to start programming in CafeOBJ, the main target are those who already have acquired a certain level of fluency, but are in need for a reference of the language.
CafeOBJ is an algebraic specification and verification language. Although it can be employed for all kind of programming (since it is Turing complete), the main target are algebraic specification of software systems. This includes programs, protocols, and all kind of interaction specifications. In addition to being a specification language, it is also a verification language, that is, a specification given in CafeOBJ can be verified within the same language environment.
Specification here means that we are trying to describe the inner workings of a software system in a mathematical way, while verification means that we give a mathematical proof of certain properties. A specification is a text, usually of formal syntax. It denotes an algebraic system constructed out of sorts (or data types) and sorted (or typed) operators. The system is characterize by the axioms in the specification. An axiom was traditionally a plain equation (``essentially algebraic''), but is now construed much more broadly. For example, CafeOBJ accommodates conditional equations, directed transitions, and (limited) use of disequality.
CafeOBJ is based on three extensions to the basic many-sorted equational logic:
There is a wide range of literature on all of these subjects for the interested reader in search for theoretical background. We refer the reader to as a starting point.
Let us start with a simple definition of a module, which are the basic building blocks of any CafeOBJ program:
mod NATPAIR {
pr(NAT)
[Pair]
var P : Pair
op <_,_> : Nat Nat -> Pair {constr}
op fst : Pair -> Nat
op snd : Pair -> Nat
eq fst( < A:Nat , B:Nat > ) = A .
eq snd( < A:Nat , B:Nat > ) = B .
}
This example already presents most of the core concepts of CafeOBJ:
pr(NAT)
[Pair]
Let us start with sorts, as they are the fundamental types.
Most programming languages allow for different sorts, or types of objects. In this respect CafeOBJ is not different and allows to have arbitrary sorts. In addition, these sorts can be ordered, more specific one sort can be declared a sub-sort of another. In the above example
[ Pair ]
a new sort called Pair
is introduced. This is a completely new sort and is in no sub-sort relation to any other sort. This is a very common case, and reflects the different types of objects in other programming languages.
In case one wants to introduce ordering in the sorts, the order can be expressed together with the definition of the sort, as in:
[ Nat < Set ]
which would introduce a new sort Set
and declares it as supersort of the (builtin) sort Nat
.
For more details concerning sorts, see sort declaration
.
CafeOBJ allows for importing and reusing of already defined modules:
pr(NAT)
for example pulls in the natural numbers (in a very minimal implementation). There are several modes of pulling in other modules, differing in the way the (semantic) models of the included module are treated.
After a statement of import, the sorts, variables, and operators of the imported modules can be used.
For more details see protecting
, extending
, using
, including
While sorts define data types, variables hold objects of a specific type, and operators define functionality. For each variable its sort has to be declared, and for each operator the signature, i.e., the sorts of the input data and the sort of the output, has to be given.
var P : Pair
op fst : Pair -> Nat
This example declares a variable P
of type pair, and an operator fst
which maps the sort Pair
to the sort Nat
, or in other words, a function that maps pairs of natural numbers to natural numbers.
We have seen already a different way to specify operators, namely
op <_,_> : Nat Nat -> Pair {constr}
which introduces an infix operator. CafeOBJ is very flexible and allows to freely specify the syntax. In an operator declaration as the above, the underscores _
represent arguments to the operator. That also means that the number of underscores must match the number of sorts given before the ->
. After the above declaration CafeOBJ will be able to parse terms like < 3 , 4 >
and correctly type them as pair.
For further details, see var
, op
.
Using sorts, variables, and operators we have specified the terms that we want to speak about. In the following equations, or sometimes called axioms, will equate different terms. Equating here is meant in the algebraic sense, but also in the term-rewriting sense, as equations form the basis of rewrite rules which provide CafeOBJ with the executable semantics:
eq fst( < A:Nat , B:Nat > ) = A .
eq snd( < A:Nat , B:Nat > ) = B .
As soon as an operator like fst
has been declared, we can give equations. In this case we define fst
of a pair to return the first element.
For further details see eq
.
In the following chapter we will include the specification of a protocol with the full code, explaining some concepts on the way.
In the following we will model a very simple protocol for cloud synchronization of a set of PCs. The full code of the actual specification, as well as parts of the verification proof score will be included and discussed.
Besides giving an example of a specification and verification, we also try to explain several of the most important concepts in CafeOBJ using rather simple examples.
One cloud computer and arbitrary many PCs have one value each that they want to keep in sync. This value is a natural number, and higher values mean more recent (like SVN revision numbers).
The Cloud can be in two states, idle and busy, while the PCs can be on of the following three states: idle, gotvalue, updated. The Cloud as well as all PCs are initially in the idle state. When a PC connects to the cloud, three things happen:
In the gotvalue state the PC compares his own value against the value it got from the cloud, and updates accordingly (changes either the cloud or the own value to the larger one). After this the PC changes into the updated state.
From the update state both the Cloud and the PC return into the idle state.
TODO include a graphic that shows this TODO
We will now go through the full specification with explanations of some of the points surfacing. We are starting with two modules that specify the possible states the cloud and the PCs can be in:
mod! CLLABEL {
[ClLabelLt < ClLabel]
ops idlecl busy : -> ClLabelLt {constr} .
eq (L1:ClLabelLt = L2:ClLabelLt) = (L1 == L2) .
}
mod! PCLABEL {
[PcLabelLt < PcLabel]
ops idlepc gotvalue updated : -> PcLabelLt {constr} .
eq (L1:PcLabelLt = L2:PcLabelLt) = (L1 == L2) .
}
Both modules define two new sorts each, the actual label, and literals for the labels. One can see that we declare the signatures of the literal labels with the ops
keyword, which introduces several operators of the same signature at the same time.
The last equation in each models provides a definition of equality by using the behavioral equality ==
. The predicate ==
is the equivalence predicate defined via reduction. Thus, the two axioms given above state that two literals for labels are the same if they are syntactically the same, since they cannot be rewritten anymore.
Furthermore, note that we choose different names for the idle state of the PCs and the cloud, to have easy separation.
The next module introduces a parametrized pair module. Parametrizing modules is a very powerful construction, and common in object oriented programming languages. In principle we leave open what are the actual components of the pairs, and only specify the operational behavior on a single pair.
In this and the next example of the multi-set, there are no additional requirements on the sorts that can be used to instantiate a pair (or multi-set). In a more general setting the argument after the double colon ::
refers to a sort, and an instantiation must be adequate for this sort (details require deeper understanding of homomorphism).
mod! PAIR(X :: TRIV,Y :: TRIV) {
[Pair]
op <_,_> : Elt.X Elt.Y -> Pair {constr}
op fst : Pair -> Elt.X
op snd : Pair -> Elt.Y
eq fst(< A:Elt.X,B:Elt.Y >) = A .
eq snd(< A:Elt.X,B:Elt.Y >) = B .
}
The next module is also parametrized, axiomatizing the concept of multi-set where a certain element can appear multiple times in the multi-set. We want to use this module to present another feature, namely the option to specify additional properties of some operators. In this case we are specifying that the constructor for sets is associative assoc
, commutative comm
, and has as identity the empty
set.
While it is easily possible to add associativity and commutativity as axioms directly, this is not advisable, especially for commutativity. Assume adding the simple equation eq A * B = B * A .
. This defines a rewrite rule from left to right. But since A
and B
are variables the can be instantiated with arbitrary subterms, and one would end up with an infinite rewriting.
mod MULTISET(X :: TRIV) {
[ Elt.X < MultiSet ]
op empty : -> MultiSet {constr} .
-- associative and commutative set constructor with identity empty
op (_ _) : MultiSet MultiSet -> MultiSet { constr assoc comm id: empty }
}
With all this set up we can defined the cloud state as a pair of a natural number, and a state. Here we see how a parametrized module is instantiated. The details of the renaming for the second element are a bit involved, but thinking about renaming of sorts and operators to match the ones given is the best idea.
Having this in mind we see that when we put the CLLABEL
into the second part of the pair, we tell the system that it should use the ClLabel
sort for the instantiation of the Elt
sort, and not the ClLabelLt
sort.
Furthermore, after the instantiation we rename the final outcome again. In this case we rename the Pair
to ClState
, and the operators to their cousins with extension in the name.
mod! CLSTATE {
pr(PAIR(NAT, CLLABEL{sort Elt -> ClLabel})*
{sort Pair -> ClState, op fst -> fst.clstate, op snd -> snd.clstate })
}
The PC state is now very similar, only that we have to have a triple (3TUPLE
is a builtin predicate of CafeOBJ), since we need one additional place for the temporary value. In the same way as above we rename the Elt
to PcLabel
and the outcome back to PcState
.
mod! PCSTATE {
pr(3TUPLE(NAT, NAT, PCLABEL{sort Elt -> PcLabel})*{sort 3Tuple -> PcState})
}
As we will have an arbitrary set of PCs, we define the multi-set of all PC states, by instatiating the multi-set from above with the just defined PcState
sort, and rename the result to PcStates
.
mod! PCSTATES {
pr(MULTISET(PCSTATE{sort Elt -> PcState})*{sort MultiSet -> PcStates})
}
Finally, the state of the whole system is declared as a pair of the cloud state and the pc states.
mod! STATE {
pr(PAIR(CLSTATE{sort Elt -> ClState},
PCSTATES{sort Elt -> PcStates})*{sort Pair -> State})
}
The final part is to specify transitions. We have described the protocol by a state machine, and the following transitions will model the transitions in this machine.
The first transition is the initialization of the synchronization by reading the cloud value, saving it into the local register, and both partners go into busy state.
Note that, since we have declared multi-set as commutative and associative, we can assume that the first element of the multi-set is actually the one we are acting on.
Transitions are different from axioms in the sense that the do not state that two terms are the same, but only that one terms can change into another.
mod! GETVALUE { pr(STATE)
trans[getvalue]:
< < ClVal:Nat , idlecl > ,
( << PcVal:Nat ; OldClVal:Nat ; idlepc >> S:PcStates ) >
=>
< < ClVal , busy > , ( << PcVal ; ClVal ; gotvalue >> S ) > .
}
The next transition is the critical part, the update of the side with the lower value. Here we are using the built-in if ... then ... else ... fi
operator.
mod! UPDATE { pr(STATE)
trans[update]:
< < ClVal:Nat , busy > ,
( << PcVal:Nat ; GotClVal:Nat ; gotvalue >> S:PcStates ) >
=>
if PcVal <= GotClVal then
< < ClVal , busy > , ( << GotClVal ; GotClVal ; updated >> S ) >
else
< < PcVal , busy > , ( << PcVal ; PcVal ; updated >> S ) >
fi .
}
The last transition is sending the both sides of the synchronization into the idle states.
mod! GOTOIDLE { pr(STATE)
trans[gotoidle]:
< < ClVal:Nat , busy > ,
( << PcVal:Nat ; OldClVal:Nat ; updated >> S:PcStates ) >
=>
< < ClVal , idlecl > , ( << PcVal ; OldClVal ; idlepc >> S ) > .
}
This completes the complete specification of the protocol, and we are defining a module CLOUD
that collects all that.
mod! CLOUD { pr(GETVALUE + UPDATE + GOTOIDLE) }
Aim of the verification is to show correctness in the sense that no two PCs are at the same time in the busy state. The idea of the proof is to show using induction on the length of transition sequences from initial states to reachable states, that for all reachable states this property is fulfilled.
More specific, we give a characterization of initial states, and show that for initial states the property holds (base case of the induction). Then we show that for all possible transitions, if the target property holds at the beginning of the transition, it also holds at the end of the transition.
Combining this with a (meta-level) induction proof on the length of transition sequences, we show that the target property holds for all reachable states.
Like with loop invariants in other verification schemes, it turns out that a single target property, the exclusion property mentioned above, does not suffice to hold over transitions, i.e., act as transition invariant. Thus, we have to extended it with additional properties.
The first part of this mini-tutorial on the specification of CloudSync contained the full code, but in the following we will, due to space reasons, only include partial code. The latest version of the CloudSync code can be obtained from .
But let us start with the definition of predicates for the initial states. The first step is to define some elementary functions on states, counting how many PCs are in a certain state:
mod! STATEfuncs {
pr(NAT + STATE)
-- no pc in gotvalue state
pred zero-gotvalue : State .
pred zero-updated : State .
...
}
We are collecting a set of predicates, indicated by their predicate name, and define apply
as an operator that checks each single predicate against a state, and forms the conjunct of the results.
mod! APPLYPREDS {
pr(STATE)
[PredName < PredNameSeq]
op (_ _) : PredNameSeq PredNameSeq -> PredNameSeq {assoc} .
op apply : PredNameSeq State -> Bool .
eq apply(P:PredName PS:PredNameSeq, S:State) = apply(P,S) and apply(PS,S) .
}
Characterization of the initial state is easy, as it only requires that all PCs as well as the cloud is in idle state.
mod! INITPREDS {
...
op cl-is-idle-name : -> PredName .
op pcs-are-idle-name : -> PredName .
...
}
In the following we define the predicate specifying initial states:
mod! INITIALSTATE {
pr(INITPREDS)
op init-name : -> PredNameSeq .
eq init-name = cl-is-idle-name pcs-are-idle-name .
pred init : State .
eq init(S:State) = apply(init-name, S) .
}
Let us now turn to the most difficult part, that is finding an invariant. This is not a one-shot technique, but mostly iterative. One starts with a set of predicates, and realizes that the proofs don't work out properly, due to some missing properties. Thus, we add new predicates and iterate until the induction proof finally succeeds.
In the following case we ended up with five different predicates that combined worked as invariant:
cloud-idle-pcs-idle
pc-clval
gotvalue
or updated
states agree.
one-active
gotvalue-cloud-value
gotvalue
state, then the value saved in the temporary storage area and the one of the cloud agree.
goal
updated
state, then the value of the PC and the value of the cloud agree.
See the mentioned web-page for the full code of these modules.
In addition to the necessity to introduce additional predicates to obtain an invariant, it also often turns out that some properties, or lemmas, have to be stated or proven so that the verification can work out. In our case some properties on if_then_else_fi
constructs, as well as consequences of rewriting are included in a module NECESSARYFACTS
.
The final - and one of the most important parts - is the proof of the two properties:
red init(S) implies invariant(S) .
red inv-condition(S, SS) .
In both cases we cannot work with a general variable S
, as in this case no rewriting can take place, and we will not obtain true. What has to be done is to provide a covering set of state expressions, i.e., a set of terms such that every possible instance of a state is also an instance of one of these terms. In our case this is quite easy to provide and consists of six different state terms, combining the three possibilities for a PC with two options of states for the cloud:
ops s1 s2 s3 s4 t1 t2 t3 t4 : -> State .
eq s1 = < < N , idlecl > , ( << M ; K ; idlepc >> PCS ) > .
eq s2 = < < N , idlecl > , ( << M ; K ; gotvalue >> PCS ) > .
eq s3 = < < N , idlecl > , ( << M ; K ; updated >> PCS ) > .
eq t1 = < < N , busy > , ( << M ; K ; idlepc >> PCS ) > .
eq t2 = < < N , busy > , ( << M ; K ; gotvalue >> PCS ) > .
eq t3 = < < N , busy > , ( << M ; K ; updated >> PCS ) > .
It is easy to see that any arbitrary state term can be obtained as instance of one of these six state terms.
What we then show is that the above properties do hold for each of these terms, and thus for each of the reachable states. In details, we show that:
red init(s1) implies invariant(s1) .
red init(s2) implies invariant(s2) .
red init(s3) implies invariant(s3) .
red init(t1) implies invariant(t1) .
red init(t2) implies invariant(t2) .
red init(t3) implies invariant(t3) .
all of these expressions reduce to true
. And furthermore, all of the following expressions, too:
red inv-condition(s1, SS) .
red inv-condition(s2, SS) .
red inv-condition(s3, SS) .
red inv-condition(t1, SS) .
red inv-condition(t2, SS) .
red inv-condition(t3, SS) .
Unfortunately, in the case of t2
this didn't turn out to be directly possible, and a further case distinction was necessary to complete the proof.
This concludes the presentation of the CloudSync protocol. We described the cloud protocol using a state system and transitions. This is just one way of implementation. There are other approaches to specification using purely term-based expressions that do not use transitions, but equational theory only. One of the strength of CafeOBJ is that it does not require any specific approach to modeling, but allows for freedom in choosing methodology.
This chapter presents all syntactic elements of CafeOBJ as well as several meta-concepts in alphabetic order. Concepts are cross-linked for easy accessibility.
Terminates the input and exit from the interpreter.
! <command>
On Unix only, forks a shell and executes the given <command>
.
#define <pattern> := <term> .
Defines
**
, **>
Starts a comment which extends to the end of the line. With the additional >
the comment is displayed while evaluated by the interpreter.
--
, -->
Starts a comment which extends to the end of the line. With the additional >
the comment is displayed while evaluated by the interpreter.
.
Input separator
=
The syntax element =
introduces an axiom of the equational theory, and is different from ==
which specifies an equality based on rewriting.
=(n)=>
, =(n,m)=>
, =()=>
=*=
The predicate for behavioral equivalence, written =*=
, is a binary operator defined on each hidden sort.
=/=
Negation of the predicate ==
.
Related: ==
==
The predicate ==
is a binary operator defined for each visible sort and is defined in terms of evaluation. That is, for ground terms t
and t'
of the same sort, t == t'
evaluates to true
iff terms reduce to a common term. This is different from the equational =
which specifies the equality of the theory.
==>
This binary predicate is defined on each visible sort, and defines the transition relation, which is reflexive, transitive, and closed under operator application. It expresses the fact that two states (terms) are connected via transitions.
Related: search predicates, trans
? [<term>]
Without any argument, shows the brief guide of online help system. With argument gives the reference manual description of term
. In addition to this, many commands allow for passing ?
as argument to obtain further help.
In case examples are provided for the <term>
, they can be displayed using ?ex <term>
. In this case the normal help output will also contain an informational message that examples are available.
When called as ?? both documentation and examples are shown.
?apropos <term> [<term> ...]
Searches all available online docs for the terms passed. Terms are separated by white space. Each term is tested independently and all terms have to match. Testing is done either by simple sub-string search, or, if the term looks like a regular expression (Perl style), by regex matching. In case a regex-like term cannot be parsed as regular expression, it is used in normal sub-string search mode.
Note: Fancy quoting with single and double quotes might lead to unexpected problems.
CafeOBJ> ?ap prec oper
will search for all entries that contain both prec
and oper
as sub-strings. Matching is done as simple sub-string match.
CafeOBJ> ?ap foo att[er]
will search for entries that contain the string foo
as well as either the string atte
or attr
.
?com [ <term> ]
List commands or declarations categorized by the key
[
Starts a sort declaration. See sort declaration for details.
accept =*= proof
switchaccept system's automatic proof of congruency of =*=
all axioms
switchControls whether axioms from included modules are shown during a show
invocation.
Related: show
always memo
switchTurns on memorization of computation also for operators without the memo
operator attribute.
Related: operator attributes, memo
:apply (<tactic> ...) [to <goal-name>]
Apply the list of tactics given within parenthesis to either the current goal, or the goal given as <goal-name>
.
Related: citp
apply <action> [ <subst> ] <range> <selection>
Applies one of the following actions reduce
, exec
, print
, or a rewrite rule to the term in focus.
reduce
, exec
, print
<range>
and <selection>
.
in this case a rewrite rule spec has to be given in the following form:
[+|-][<mod_name>].<rule-id>
where <mod_name>
is the name of a module, and <rule-id>
either a number n - in which case the n. equation in the current module is used, or the label of an equation. If the <mod_name>
is not given, the equations of the current module are considered. If the leading +
or no leading character is given, the equation is applied left-to-right, which with a leading -
the equation is applied right-to-left.
The <subst>
is of the form
with { <var_name> = <term> } +,
and is used when applying a rewrite rule. In this case the variables in the rule are bound to the given term.
<range>
is either within
or at
. In the former case the action is applied at or inside the (sub)term specified by the following selection. In the later case it means exactly at the (sub)term.
Finally, the <selection>
is an expression
<selector> { of <selector> } *
where each <selector>
is one of
top
, term
subterm
choose
)
( <number_list> )
(2 1)
indicates a subterm by tree search. (2 1)
means the first argument of the second argument.
[ <number1> .. <number2> ]
This selector can only be used with associative operators. It indicates a subterm in a flattened structure and selects the subterm between and including the two numbers given. [n .. n]
can be abbreviated to [n]
.
Example: If the term is a * b * c * d * e
, then the expression [2 .. 4]
selects the subterm b * c * d
.
{ <number_set> }
Example: If the operator _*_
is declared as associative and commutative, and the current term is b * c * d * c * e
, then then the expression {2, 4, 5}
selects the subterm c * c * e
.
:auto
Applies the following set of tactics: (SI CA TC IP RD)
.
Related: citp
auto context
switchPossible values: on
or off
, default is off
.
If this switch is on
, the context will automatically switch to the most recent module, i.e., defining a module or inspecting a module's content will switch the current module.
autoload <module-name> <file-name>
When evaluating a
Related: no autoload
ax [ <label-exp> ] <term> = <term>
.(pignose)
axioms { <decls> }
Block enclosing declarations of variables, equations, and transitions. Other statements are not allowed within the axioms
block. Optional structuring of the statements in a module.
Related: trans
, eq
, var
, imports
, signature
:backward equation|rule
Like :equation
and :rule
, but exchange the left and right side.
Related: :rule
, :equation
, :cp
, citp
bax [ <label-exp> ] <term> = <term>
.(pignose)
bceq [ <label-exp> ] <term> = <term> if <boolterm> .
Defines a behavioral conditional equation. For details see ceq
.
bcrule [ <label-exp> ] <term> => <term> if <term> .
Synonym of bctrans
Related: bctrans
bctrans [ <label-exp> ] <term> => <term> if <bool> .
Defines a behavioral conditional transition. For details see ctrans
.
Related: btrans
, ctrans
, trans
beq [ <label-exp> ] <term> = <term> .
Defines a behavioral equation. For details see eq
.
bgoal <term> .
(pignose)
{bguess | :bguess} {imply|and|or} [ with <predicate name> ]
Try to find true/false assignments which satisfies the Bool term specified by 'binspect' or ':binspect'.
binspect [in <module-name> :] <boolean-term> .
Start an inspection of a Boolean term, that is, and abstracted form of the Boolean term is constructed. The abstracted term is shown (like calling bshow
.
CafeOBJ> module BTE { [S]
preds p1 p2 p3 p4 p5 p6 p7 : S
ops a b c : -> S .
}
CafeOBJ> binspect in BTE : (p1(X:S) or p2(X)) and p3(Y:S) or (p4(Y) and p1(Y)) .
...
--> ((p4(Y:S) and p1(Y)) xor ((p3(Y) and p1(X:S)) xor ((p2(X) and (p3(Y) and p1(X))) xor ((p3(Y) and p2(X)) xor ((p3(Y) and (p2(X) and (p4(Y) and p1(Y)))) xor ((p3(Y) and (p2(X) and (p1(X) and (p1(Y) and p4(Y))))) xor (p1(X) and (p3(Y) and (p1(Y) and p4(Y))))))))))
...
:binspect [in <goal-name> :] <boolean-term> .
Used during CITP proofs instead of binspect
bop <op-spec> : <sorts> -> <sort>
Defines a behavioral operator by its domain, co-domain, and the term construct. <sorts>
is a space separated list of sort names containing exactly one hidden sort. <sort>
is a single sort name.
For <op-spec>
see the explanations of op
.
Related: op
bpred <op-spec> : <sorts>
Short hand for op <op-spec> : <sorts> -> Bool
defining a behavioral predicate.
breduce [ in <mod-exp> : ] <term> .
Reduce the given term in the given module, if <mod-exp>
is given, otherwise in the current module.
For breduce
equations, possibly conditional, possibly behavioral, are taken into account for reduction.
{bresolve | :bresolve} [<limit>] [all]
Computes all possible variable assignments that render an abstracted term true
. The variant with leading colon is for usage during a CITP proof. If an optional argument 'all' is specified, all solutions will be searched. Optional
CafeOBJ> bresolve 2 all
** (1) The following assignment(s) makes the term to be 'true'.
[1] { P-3:Bool |-> true }
where
p-3 = P4(Y:S)
[2] { P-4:Bool |-> true }
where
p-4 = P1(X:S)
** (2) The following assignment(s) makes the term to be 'true'.
[1] { P-1:Bool |-> true, P-2:Bool |-> true }
where
p-1 = P3(Y:S)
p-2 = P2(X:S)
...
brule [ <label-exp> ] <term> => <term> .
Synonym of btrans
.
Related: btrans
{bshow | :bshow} [{ tree | grind }]
Shows the abstracted Boolean term computed by binspect
. If the argument tree
is given, prints out a the abstracted term in tree form. The variant with leading colon is for usage during a CITP proof.
CafeOBJ> bshow
((P-1:Bool and (P-2:Bool and (P-3:Bool and P-4:Bool))) xor ((P-1 and (P-2 and (P-4 and (P-5:Bool and P-3)))) xor ((P-2 and (P-1 and (P-5 and P-3))) xor ((P-5 and P-3) xor ((P-4 and (P-3 and P-5)) xor ((P-4 and P-3) xor (P-2 and P-1)))))))
where
P-1:Bool |-> p4(Y:S)
P-2:Bool |-> p1(Y:S)
P-3:Bool |-> p3(Y:S)
P-4:Bool |-> p1(X:S)
P-5:Bool |-> p2(X:S)
bsort token-predicate creater printer term-predicate
Defines a built-in sort. Internal use only.
btrans [ <label-exp> ] <term> => <term> .
Defines a behavioral transition. For details see trans
.
Related: bctrans
, ctrans
, trans
cbred [ in <mod-exp> :] <term> .
circular coinductive reduction: see Goguen, Lin, Rosu: Circular Coinductive Rewriting (Proceedings of Automated Software Engineering 2000) for details.
cd <dirname>
Change the current working directory, like the Unix counterpart. The argument is necessary. No kind of expansion or substitution is done.
ceq [ <label-exp> ] <term> = <term> if <boolterm> .
Defines a conditional equation. Spaces around the if
are obligatory. <boolterm>
needs to be a Boolean term. For other requirements see eq
.
check <options>
This command allows for checking of certain properties of modules and operators.
check regularity <mod_exp>
<mod_exp>
is regular.
check compatibility <mod_exp>
<mod_exp>
is compatible, i.e., every application of every rewrite rule to every well-formed term results in a well-formed term. (This is not necessarily the case in order-sorted rewriting!)
check laziness <op_name>
<op_name>
is given, all operators of the current module are checked.
Related: regularize
check <something>
switchThese switches turn on automatic checking of certain properties:
check coherency
check compatibility
check
command
check import
check regularity
check
command
check sensible
choose <selection>
Chooses a subterm by the given <selection>
. See apply
for details on <selection>
.
Related: strat
in operator attributes, start
, apply
Constructor Based Induction Theorem Prover
The sub-system provides a certain level of automatization for theorem proving.
Please see the accompanying manual for CITP for details.
Related: :attr
, :reset
, :embed
, :use
, :ord
, :imp
, :def
, :ctf-
, :ctf
, :csp-
, :csp
, :red
, :select
, :backward
, :rule
, :equation
, :cp
, :init
, :roll
, :auto
, :ind
, :apply
, :goal
clause <term> .
(pignose)
clean memo
Resets (clears) the memo storage of the system. Memorized computations are forgotten.
Related: clean memo switch
clean memo
switchPossible values: on
, off
, default off
.
tells the system to be forgetful.
close
This command closes a modification of a module started by open
.
Related: open
commands
Print outs the list of main toplevel commands.
The interpreter accepts the following strings as start of a comment that extends to the end of the line: --
, -->
, **
, **>
.
The difference in the variants with >
is that the comment is displayed when run through the interpreter.
cond limit
switchSetting maximal number of evaluation of condition part of an axiom. This is useful for detecting a kind of inifinite loop of rewriting.
cont
In step mode, continues the reduction until a stop pattern has been found.
:cp { "[" <label> "]" | "(" <sentence> . ")" } >< { "[" <label> "]" | "(" <sentence> .")" }
Computes the critical pair of the two given equations. Here either a label or a full equation can be used to specify the equations.
Related: citp
:cp (ceq top(sq(S@Sys)) = I@Pid if pc(S@Sys,I@Pid) = cs .)
><
(ceq top(sq(S@Sys)) = J@Pid if pc(S@Sys,J@Pid) = cs .)
crule [ <label-exp> ] <term> => <term> if <term> .
Synonym of ctrans
:csp { eq [ <label-exp>] <term> = <term> . ...}
Applies case splitting after a set of equations. Each of these equations creates one new sub-goal with the equation added.
The system does not check whether given set of equations exhausts all possible values.
Not discharged sub-goals will remain in the reduced form.
:csp- { eq [ <label-exp>] <term> = <term> . ...}
Like :csp
, but if sub-goals are not discharged, the CITP prover returns to the original state before the reduce action.
:ctf { eq [ <label-exp> ] <term> = <term> .}
Applies case splitting after a set of boolean expressions. Not discharged sub-goals will remain in the reduced form.
:ctf- { eq [ <label-exp> ] <term> = <term> .}
Like :ctf
, but if sub-goals are not discharged, the CITP prover returns to the original state before the reduce action.
ctrans [ <label-exp> ] <term> => <term> if <term> .
Defines a conditional transition. For details see trans
and ceq
.
Related: bctrans
, btrans
, trans
db reset
(pignose)
:def <symbol> = { <ctf> | <csp> | <init> }
Assigns a name to a specific case splitting (:ctf
or :csp
) or induction :ind
), so that it can be used as tactics in :apply
.
Related: citp
:def name-0 = :ind { :on (<Variable>...) :base <Term> . :step <Term> . }
:def name-1 = :ctf [ <Term> . ]
:def name-2 = :ctf-{ eq LHS = RHS . }
:def name-3 = :csp { eq lhs1 = rhs1 . eq lhs2 = rhs2 . }
:def name-4 = :csp-{ eq lhs3 = rhs3 . eq lhs4 = rhs4 . }
:apply(name-0 TC name-1 name-2 name-3 name-4)
demod
(pignose)
:describe proof
Describes the current proof in more detail.
PNAT> :describe proof
==> root*
-- context module: #Goal-root
-- targeted sentences:
eq [lemma-1]: M:PNat + 0 = M .
eq [lemma-2]: M:PNat + s N:PNat = s (M + N) .
[si] 1*
-- context module: #Goal-1
-- targeted sentences:
eq [lemma-1]: 0 + 0 = 0 .
eq [lemma-2]: 0 + s N:PNat = s (0 + N) .
...
describe <something>
Similar to the show
command but with more details. Call describe ?
for the possible set of invocations.
Related: show
dirs
Displays the current push stack.
Related: popd
, pwd
, pushd
, cd
, ls
dribble { <file-name> | .}
If
:embed (<label> ... <label>) as <module_name>
Incorporate proved goals into the module specified by
eof
Terminates reading of the current file. Allows for keeping untested code or documentations below the eof
mark. Has to be on a line by itself without leading spaces.
eq [ <label-exp> ] <term> = <term> .
Declares an axiom, or equation.
Spaces around the =
are necessary to separate the left from the right hand side. The terms given must belong to the same connected component in the graph defined by the sort ordering.
In simple words, the objects determined by the terms must be interpretable as of the same sort.
The optional part <label-exp>
serves two purposes, one is to give an axiom an identifier, and one is to modify its behavior. The <label-exp>
is of the form:
[ <modifier> <label> ] :
Warning: The square brackets here are not specifying optional components, but syntactical elements. Thus, a labeled axiom can look like:
eq[foobar] : foo = bar .
The <modifier>
part is used to change the rewriting behavior of the axiom. There are at the moment two possible modifiers, namely :m-and (:m-and-also)
and :m-or (:m-or-else)
. Both make sense only for operators where the arguments come from an associative sort. In this case both modifiers create all possible permutations of the arguments and rewrite the original term to the conjunction in case of :m-and
or to the disjunction in case of :m-or
of all the generated terms.
Assume that NatSet
is a sort with associative constructor modeling a set of natural number, and let
pred p1: Nat .
ops q1 q2 : NatSet -> Bool .
eq [:m-and]: q1(N1:Nat NS:NatSet) = p1(N1) .
eq [:m-or]: q2(N1:Nat NS:NatSet) = p1(N1) .
In this case an expression like q1(1 2 3)
would reduce to p1(1) and p1(2) and p1(3)
(modulo AC), and q2(1 2 3)
into the same term with or
instead.
:equation
Adds the critical pair computed by the last :cp
command as equation to the current goal.
esc return
In case that, after hitting return
expecting some feed-back, no such feed-back whatsoever is returned, typing the escape
key followed by the return
key will make the interpreter discard the preceding input and make a fresh start.
exec limit
switchPossible values: integers, default limit 4611686018427387903.
Controls the number of maximal transition steps.
Related: reduce
exec trace
switchPossible values: on
off, default
off`.
controls whether further output is provided during reductions.
Related: reduce
exec! [ in <mod-exp> : ] <term> .
Obsolete command. Implicitly invokes RWL search predicate in a specific manner.
execute [ in <mod-exp> : ] <term> .
Reduce the given term in the given module, if <mod-exp>
is given, otherwise in the current module.
For execute
equations and transitions, possibly conditional, are taken into account for reduction.
extending ( <modexp> )
Imports the object specified by modexp
into the current module, allowing models to be inflated, but not collapsing. See module expression
for format of modexp
.
Related: using
, protecting
, including
find {+rule | -rule}
Find all axioms which possibly rewrite the current term.
find all rules
switchIf this switch is on, the apply
command will search for applicable rules not only in the set of user-defined equations, but also in those added by the system.
flag(<name>, { on | off })
(pignose)
full reset
Reinitializes the internal state of the system. All supplied modules definitions are lost.
gendoc <pathname>
generates reference manual from system's on line help documents, and save it to pathname
.
:goal { <sentence> . ... }
Define the initial goal for CITP
Related: citp
CafeOBJ> select PNAT .
PNAT> :goal {
eq [lemma-1]: M:PNat + 0 = M .
eq [lemma-2]: M:PNat + s N:PNat = s( M + N ) .
}
goal <term> .
(pignose)
:imp "[" <label> "]" by "{" <variable> <- <term>; ..."}"
TODO (future extension)
Related: citp
imports { <import-decl> }
Block enclosing import of other modules (protecting
etc). Other statements are not allowed within the imports
block. Optional structuring of the statements in a module.
Related: using
, protecting
, including
, extending
, axioms
, signature
include BOOL
switchPossible values: on
off
, default on
.
By default a couple of built-in modules are implicitly imported with protecting mode. In particular, BOOL is of practical importance. It defines Boolean operators. It is imported to admit conditional axioms.
This switch allows to disable automatic inclusion of BOOL.
include RWL
switchPossible values: on
off
, default off
.
This switch allows to disable automatic inclusion of RWL.
including ( <modexp> )
Imports the object specified by modexp
into the current module.
See module expression
for format of modexp
.
Related: module expression, using
, protecting
, extending
:ind { on (<variable> ...) | '{' on (<variable> ...) base (<Term> . ... <Term> .) step (<Term> . ... <Term> .) '}'
':ind on (
Related: citp
:ind on (M:PNat)
:ind { on (M:PNat)
base (<Term> . ... <Term> .)
step (<Term> . ... <Term> .)
}
init [as <name>] { "[" <label> "]" | "(" <sentence> "")} by "{" <variable> <- <term>; ... "}"
Instantiates an equation specified by <label>
by replacing the <variable>
s in the equation with the respective <term>
s. The resulting equation is added to the set of axioms. If optional as <name>
is given, label of the instantiated axiom is overwritten by
Related: open
:init [as <name>] { "[" <label> "]" | "(" <sentence> "")} by "{" <variable> <- <term>; ... "}"
Instantiates an equation specified by <label>
by replacing the <variable>
s in the equation with the respective <term>
s. The resulting equation is added to the set of axioms. If optional as <name>
is given, label of the instantiated axiom is overwritten by
Related: citp
input <pathname>
Requests the system to read the file specified by the pathname. The file itself may contain input
commands. CafeOBJ reads the file up to the end, or until it encounters a line that only contains (the literal) eof
.
inspect <term>
Inspect the internal structure of <term>
.
Parameterized modules allow for instantiation. The process of instantiation binds actual parameters to formal parameters. The result of an instantiation is a new module, obtained by replacing occurrences of parameter sorts and operators by their actual counterparts. If, as a result of instantiation, a module is imported twice, it is assumed to be imported once and shared throughout.
Instantiation is done by
<module_name> ( <bindings> )
where <module_name>
is the name of a parameterized module, and <bindings>
is a comma-separated list of binding constructs.
you may bind an already declared view to a parameter:
<parameter> <= <view_name>
If a module M
has a parameter X :: T
and a view V
from T
to M'
is declared, V
may be bound to X
, with the effect that
The sort and operator names of T
that appear in the body of M
are replaced by those in M'
, in accordance with V
,
The common submodules of M
and M'
are shared.
In this case the view is declared and used at the same time.
<parameter> <= view to <mod_name> { <view_elements> }
See view
for details concerning <view_elements>
. The from
parameter in the view
declaration is taken from <parameter>
.
To make notation more succinct, parameters can be identified also by position instead of names as in
<mod_name> ( <view_name>, <view_name> )
which would bind the <view_name>
s to the respective parameters of the parameterized module <mod_name>
.
This can be combined with the ephemeral definition of a view like in the following example (assume ILIST
has two parameters):
module NAT-ILIST {
protecting ( ILIST(SIMPLE-NAT { sort Elt -> Nat },
DATATYPE { sort Elt -> Data }) )
}
:is
Boolean expression: A :is B
where A
is a term and B
is a sort. Returns true if A
is of sort B
.
let <identifier> = <term> .
Using let
one can define aliases, or context variables. Bindings are local to the current module. Variable defined with let
can be used in various commands like reduce
and parse
.
Although let
defined variable behave very similar to syntactic shorthands, they are not. The right hand side <term>
needs to be a fully parsable expression.
lex (<op>, ..., <op>)
(pignose)
libpath
switchPossible values: list of strings.
The switch libpath
contains a list of directories where CafeOBJ searches for include files. Addition and removal of directories can be done with
set libpath + <path1>:<path2>:...
set libpath - <path1>:<path2>:...
or the full libpath reset by set libpath <path1>:<path2>:...
The current directory has a privileged status: It is always searched first and cannot be suppressed.
lisp
Evaluates the following lisp expression.
CafeOBJ> lisp (+ 4 5)
(+ 4 5) -> 9
lispq
Evaluates the following lisp expression, but does not display the result (q for quiet).
list { axiom | sos | usable | flag | param | option | demod }
(pignose)
look up <something>
displays the location (module) and further information where <something>
has been defined.
open INT .
%INT> look up Nat .
Nat
- sort declared in NAT-VALUE
- operator:
op Nat : -> SortId { constr prec: 0 }
-- declared in module NAT-VALUE
%INT>
ls <pathname>
lists the given pathname
. Argument is obligatory.
make <mod_name> ( <mod_exp> )
This commands defines a new module <mod_name>
by evaluating the module expression <mod_exp>
.
Related: module expression
match <term_spec> to <pattern> .
Matches the term denoted by <term_spec>
to the pattern. <term_spec>
is either top
or term
for the term set by the start
command; subterm
for the term selected by the choose
command; it
has the same meaning as subterm
if choose
was used, otherwise the same meaning as top
, or a normal term expression.
The given <pattern>
is either rules
, -rules
, +rules
, one of these three prefixed by all
, or a term. If one of the rules
are given, all the rules where the left side (for +rules
), the right side (for -rules
), or any side (for rules
) matches. If the all
(with separating space) is given all rules in the current context, including those declared in built-in modules, are inspected.
If a term is given, then the two terms are matched, and if successful, the matching substitution is printed.
memo
switchcontrols the memorization of computations. The system memorizes evaluations of operators declared with the memo
operator attribute. Turning this switch off disables all memorization.
[sys:]module[!|*] <modname> [ ( <params> ) ] [ <principal_sort_spec> ] { mod_elements ... }
Defines a module, the basic building block of CafeOBJ. Possible elements are declarations of
protecting
, extending
, including
, using
sort declaration
var
op
, eq
, ceq
, bop
, beq
, bceq
trans
, ctrans
, btrans
, bctrans
modname
is an arbitrary string.
module*
introduces a loose semantic based module.
module!
introduces a strict semantic based module.
module
introduces a module without specified semantic type.
If params
are given, it is a parameterized module. See parameterized module
for more details.
If principal_sort_spec
is given, it has to be of the form principal-sort <sortname>
(or p-sort <sortname>
). The principal sort of the module is specified, which allows more concise view
s from single-sort modules as the sort mapping needs not be given.
module expression
In various syntax elements not only module names itself, but whole module expressions can appear. A typical example is
open <mod_exp> .
which opens a module expression. The following constructs are supported:
<mod_exp> * { <mappings> }
This expressions describes a new module where sort and/or operators are renamed. <mappings>
are like in the case of view
a comma separated list of mappings of either sorts (sort
and hsort
) or operators (op
and bop
). Source names may be qualified, while target names are not, they are required to be new names. Renaming is often used in combination with instantiation.
<mod_exp> + <mod_exp>
This expression describes a module consisting of all the module elements of the summands. If a submodule is imported more than once, it is assumed to be shared.
names <mod-exp>
.List up all the named objects in module
no autoload <module-name>
Stop autoload
of module with the name autoload
command.
Related: autoload
:normalize { on | off}
Normalize the LHS of an instance of the axiom generated by :init command.
Related: citp
Variables and constants can be declared on-the-fly (or inline). If an equation contains a qualified variable (see qualified term), i.e., <name>:<sort-name>
, then from this point on within the current equation only <name>
is declared as a variable of sort <sort-name>
.
It is allowed to redeclare a previously defined variable name via an on-the-fly declaration, but as mentioned above, not via an explicit redeclaration.
Using a predeclared variable name within an equation first as is, that is as the predeclared variable, and later on in the same equation with an on-the-fly declaration is forbidden. That is, under the assumption that A
has been declared beforehand, the following equation is not valid:
eq foo(A, A:S) = A .
On-the-fly declaration of constants are done the same way, where the <name>
is a constant name as in `a:Nat
. Using this construct is similar to defining an operator
op <name> : -> <sort>
or in the above example, op a : -> Nat .
, besides that the on-the-fly declaration of constants, like to one of variables, is only valid in the current context (i.e., term or axiom). These constant definitions are quite common in proof scores.
Related: var
op <op-spec> : <sorts> -> <sort> { <attribute-list> }
Defines an operator by its domain, co-domain, and the term construct. <sorts>
is a space separated list of sort names, <sort>
is a single sort name. <op-spec>
can be of the following forms:
the <op-spec>
does not contain a literal _
: This defines a normal prefix operator with domain <sorts>
and co-domain <sort>
Example: op f : S T -> U
mixfix-spec
the <op-spec>
contains exactly as many literal _
as there are sort names in <sorts>
: This defines an arbitrary mixfix (including postfix) operator where the arguments are inserted into the positions designated by the underbars.
Example: op _+_ : S S -> S
For the description of <attribute-list>
see the entry for operator attributes.
open <mod_exp> .
This command opens the module specified by the module expression <mod_exp>
and allows for declaration of new sorts, operators, etc.
Related: select
, module expression
, close
operator attributes
In the specification of an operator using the op
(and related) keyword, attributes of the operator can be specified. An <attribute-list>
is a space-separate list of single attribute definitions. Currently the following attributes are supported
associative
assoc
commutative
comm
itempotence
idem
id: <const>
<const>
prec: <int>
l-assoc
and r-assoc
constr
strat: ( <int-list> )
specifies the evaluation strategy. Each integer in the list refers to an argument of the operator, where 0
refers to the whole term, 1
for the first argument, etc. Evaluation proceeds in order of the <int-list>
. Example:
op if_then_else_fi : Bool Int Int -> Int { strat: (1 0) }
In this case the first argument (the Boolean term) is tried to be evaluated, and depending on that either the second or third. But if the first (Boolean) argument cannot be evaluated, no evaluation in the subterms will appear.
Using negative values allows for lazy evaluation of the corresponding arguments.
memo
memo
switch for details.
Remarks:
Several operators of the same arity/coarity can be defined by using ops
instead of op
:
ops f g : S -> S
For the case of mixfix operators the underbars have to be given and the expression surrounded by parenthesis:
ops (_+_) (_*_) : S S -> S
Spaces can be part of the operator name, thus an operator definition of op foo op : S -> S
is valid, but not advisable, as parsing needs hints.
A single underbar cannot be an operator name.
Related: bop
operator precedence
CafeOBJ allows for complete freedom of syntax, in particular infix operators and overloading. To correctly parse terms that are ambiguous, all operators have precedence values. These values can be adjusted manually during definition of the operator (see operator attributes). In absence of manual specification of the operator precedence, the values are determined by the following rules:
op f : S1 .. Sk -> S
, receive operator precedence value 0.op u_ : S1 -> S
, receive precedence 15.op _ arg-or-op _ : S1 .. Sk -> S
, receive precedence 41.a _ b
, etc.) receive precedence 0.Related: operator attributes
option { reset | = <name> }
(pignose)
:order (<op>, ..., <op>)
param(<name>, <value>)
(pignose)
parameterized module
A module with a parameter list (see module
) is a parameterized module. Parameters are given as a comma (,
) separated list. Each parameter is of the form [ <import_mode> ] <param_name> :: <module_name>
(spaces around ::
are obligatory).
The parameter's module gives minimal requirements on the module instantiation.
Within the module declaration sorts and operators of the parameter are qualified with .<parameter_name>
as seen in the example below.
Related: qualified sort
mod* C {
[A]
op add : A A -> A .
}
mod! TWICE(X :: C) {
op twice : A.X -> A.X .
eq twice(E:A.X) = add.X(E,E) .
}
parse [ in <mod-exp> : ] <term> .
Tries to parse the given term within the module specified by the module expression <mod-exp>
, or the current module if not given, and returns the parsed and qualified term.
In case of ambiguous terms, i.e., different possible parse trees, the command will prompt for one of the trees.
Related: qualified term
parse normalize
switchIf this switch is 'on' (defalult is 'off'), terms with associative operators are always parsed as right associative.
popd
Changes the current working directory to the last on on the push stack.
Related: dirs
, pwd
, pushd
, cd
, ls
pred <op-spec> : <sorts>
Short hand for op <op-spec> : <sorts> -> Bool
defining a predicate.
prelude <file>
Loads the given <file>
as prelude. That is, a call to reset
will reset the definitions made in this file.
Related: full reset
, reset
print depth
switchPossible values: natural numbers, default unlimited
.
Controls to which depth terms are printed.
print mode
switchPossible values: normal
fancy
tree
s-expr
Selects one of the print modes.
print trs
switchPossible values: on
off
, default off
If set to on
, print the rules used during reduction of =(_,_)=>+_if_suchThat_{_}
.
Related: search predicates
protect <module-name>
Protect a module from being overwritten. Some modules vital for the system are initially protected. Can be reversed with unprotect
.
Related: unprotect
protecting ( <modexp> )
Imports the object specified by modexp
into the current module, preserving all intended models as they are. See module expression
for format of modexp
.
Related: including
, using
, extending
provide <feature>
Discharges a feature requirement: once provide
d, all the subsequent require
ments of a feature are assumed to have been fulfilled already.
Related: require
pushd <directory>
Changes the working directory to <directory>
, and puts the current directory onto the push stack. Going back can be done with pop
.
Related: dirs
, pwd
, popd
, cd
, ls
pvar <var-name> : <sort-name>
(pignose)
pwd
Prints the current working directory.
Related: dirs
, popd
, pushd
, ls
, cd
CafeOBJ allows for using the same name for different sorts, operators, and parameters. One example is declaring the same sort in different modules. In case it is necessary to qualify the sort, operator, or parameter, the intended module name can be affixed after a literal .
: <name>.<modname>
Example: In case the same sort Nat
is declared in both the module SIMPLE-NAT
and PANAT
, one can use Nat.SIMPLE-NAT
to reference the sort from the former module.
Furthermore, a similar case can arise when operators of the same name have been declared with different number of arguments. During operator renaming (see view) the need for qualification of the number of parameters might arise. In this case the number can be specified after an affixed /
: <opname>/<argnr>
Related: qualified term
, parameterized module
qualified term
In case that a term can be parsed into different sort, it is possible to qualify the term to one of the possible sorts by affixing it with : <sort-name>
(spaces before and after the :
are optional).
Related: parse
(1):NzNat
(2):Nat
quiet
switchPossible values: on
off
, default off
If set to on
, the system only issues error messages.
Related: verbose
quit
Leaves the CafeOBJ interpreter.
{ :red | :exec | :bred } [in <goal-name> :] <term> .
reduce the term in specified goal
Related: citp
reduce [ in <mod-exp> : ] <term> .
Reduce the given term in the given module, if <mod-exp>
is given, otherwise in the current module.
For reduce
only equations and conditional equations are taken into account for reduction.
reduce conditions
switchPossible values: on
off
, default off
.
When using apply
to step through a reduction, this switch allows to turn on automatic reduction of conditions in conditional equations.
Related: apply
regularize <mod-name>
Regularizes the signature of the given module, ensuring that every term has exactly one minimal parse tree. In this process additional sorts are generated to ensure unique least sort of all terms.
Modules can be automatically regularized by the interpreter if the regularize signature
switch is turn to on
.
regularize signature
switchSee `regularize
require <feature> [ <pathname> ]
Requires a feature, which usually denotes a set of module definitions. Given this command, the system searches for a file named the feature, and read the file if found. If the <feature>
contains ::
, they are treated as path separators.
If a pathname is given, the system searches for a file named the pathname instead.
Related: provide
CafeOBJ> require foo::bar
would search for foo/bar.cafe
in the pathes from libpath
:reset
Discard the current proof session.
reset
Restores the definitions of built-in modules and preludes, but does not affect other modules.
Related: prelude
, full reset
resolve {. | <file-path> }
(pignose)
restore <pathname>
Restores module definitions from the designated file pathname
which has been saved with the save
command. input
can also be used but the effects might be different.
TODO -- should we keep the different effects? What is the real difference?
Related: save-system
, save
, input
rewrite limit
switchPossible values: positive integers, default not specified.
Allows limiting the number of rewrite steps during a step-wise execution.
Related: step switch
:roll back
Reverts the strategy that led to the current target goal. The current target goal is removed from the proof tree.
Related: citp
:rule
Adds the critical pair computed by the last :cp
command as rule to the current goal.
rule [ <label-exp> ] <term> => <term> .
Synonym of trans
.
Related: trans
save <pathname>
Saves module definitions into the designated file pathname
. File names should be suffixed with .bin
.
save
also saves the contents of prelude files as well as module definitions given in the current session.
Related: save-system
, restore
, input
save-option <name>
(pignose)
scase (<term>) in (<mod-exp>) as <name> { <decl> ..} : <term> .
Obsolete citp command. Split the goal by user specified cases.
search predicates
CafeOBJ provides a whole set of search predicates, that searches the reachable states starting from a given state, optionally checking additional conditions. All of them based on the following three basic ones:
S =(n,m)=>* SS [if Pred]
search states reachable by 0 or more transitions;S =(n,m)=>+ SS [if Pred]
search states reachable by 1 or more transitions;S =(n,m)=>! SS [if Pred]
search states reachable by 0 or more transitions, and require that the reached state is a final state, i.e., no further transitions can be applied.To allow for conditional transitions, a transition is only considered in the search if Pred
holds.
The parameters n
and m
in these search predicates:
n
, a natural number or *
, gives the maximal number of solutions to be searched. If *
is given all solutions are searched exhaustively.m
, a natural number but not *
, gives the maximal depth up to which search is performed.The predicates return true if there is a (chain of) transitions that fits the parameters (n
,m
, and *
, +
, !
) and connects S
with SS
.
There are two orthogonal extension to this search predicate, one adds a suchThat
clause, one adds a withStateEq
clause.
S =(n,m)=>* SS [if Pred1] suchThat Pred2
!
and +
) In this case not only the existence, of a transition sequence is tested, but also whether the predicate Pred2
, which normally takes S
and SS
as arguments, holds.
S =(n,m)=>* SS [if Pred1] withStateEq Pred2
!
and +
) Pred2
is used to determine whether a search continues at SS
or not, by comparing SS
with all states that have been traversed in the current search. If the predicate Pred2
returns true on the combination of SS
as first argument, and any of the previously visited states as second argument, then the search is not continued after SS
. (This is a kind of loop detection.)
These two cases can also be combined into
S =(n,m)=>* SS [if Pred1] suchThat Pred2 withStateEq Pred3
:select <goal-name>
Select a goal for further application of tactics.
Related: citp
select <mod_exp> .
Selects a module given by the module expression <mod_exp>
as the current module. All further operations are carried out within the given module. In contrast to open
this does not allow for modification of the module, e.g., addition of new sorts etc.
Related: module expression
, open
:set(<name>, { on | off | show })
Set or show various flags of CITP CafeOBJ.
Related: citp
set <name> [option] <value>
Depending on the type of the switch, options and value specification varies. Possible value types for switches are Boolean (on
, off
), string ("value"
), integers (5434443), lists (lisp syntax).
For a list of all available switches, use set ?
. To see the current values, use show switches
. To single out two general purpose switches, verbose
and quiet
tell the system to behave in the respective way.
:show goal|unproved|proof|discharged
Shows the current goal, the up-to-now unproven (sub-)goals, and the current proof.
PNAT> :show proof
root*
[si] 1*
[ca] 1-1*
[ca] 1-2*
[tc] 1-2-1*
[si] 2*
[ca] 2-1*
[ca] 2-2*
[tc] 2-2-1*
PNAT>
show <something>
The show
command provides various ways to inspect all kind of objects of the CafeOBJ language. For a full list call show ?
.
Some of the more important (but far from complete list) ways to call the show
command are:
show [ <modexp> ]
- describes the current modules of the one specified as argumentshow module tree [ <modexp> ]
- displays submodules of show switches
- lists all possible switchesshow term [ tree ]
- displays a term, possible in tree formatSee the entry for switches
for a full list.
show mode
switchPossible values for set show mode <mode>
are cafeobj
and meta
.
sigmatch (<mod-exp>) to (<mod-exp>)
(pignose)
signature { <sig-decl> }
Block enclosing declarations of sorts and operators. Other statements are not allowed within the signature
block. Optional structuring of the statements in a module.
Related: op
, sort
, imports
, axioms
CafeOBJ supports two kind of sorts, visible and hidden sorts. Visible sorts are introduced between [
and ]
, while hidden sorts are introduced between *[
and ]*
.
[ Nat ]
*[ Obs ]*
Several sorts can be declared at the same time, as in [ Nat Int ]
.
Since CafeOBJ is based on order sorting, sorts can form a partial order. Definition of the partial order can be interleaved by giving
[ <sorts> < <sorts> ]
Where sorts
is a list of sort names. This declaration defines an inclusion relation between each pair or left and right sorts.
[ A B , C D < A < E, B < D ]
defines five sorts A
,...,E
, with the following relations: C < A
, D < A
, A < E
, B < D
.
sos { = | + | - } { <clause> , ... }
(pignose)
:spoiler { on | off}
If the spoiler flag is on, after a strategy other than RD and SI has been applied, the generated sub-goals are automatically checked for provability using the RD strategy. Defaults to off
.
Related: citp
start <term> .
Sets the focus onto the given term <term>
of the currently opened module or context. Commands like apply
, choose
, or match
will then operate on this term.
statistics
switchPossible values: on
off
, default on
.
After each reduction details about the reduction are shown. Information shown are the time for parsing the expression, the number of rewrites and run time during rewriting, and the number of total matches performed during the reduce.
step
switchPossible values: on
off
, default off
.
With this switch turned on, rewriting proceeds in steps and prompts the user interactively. At each prompt the following commands can be given to the stepper (with our without leading colon :
):
help
h
, ?
) print out help page next
n
) go one step continue
c
) continue rewriting without stepping quit
q
) leave stepper continuing rewrite abort
a
) abort rewriting rule
r
) print out current rewrite rule subst
s
) print out substitution limit
l
) print out rewrite limit count pattern
p
) print out stop pattern stop [<term>] .
rwt [<number>] .
Other standard CafeOBJ commands that can be used are show
, describe
, dirs
, set
, cd
, ls
, pwd
, pushd
, popd
, lisp
, lispq
, and (on Unix only) !
.
stop
Equivalent to stop pattern switch
stop pattern
switchIn step mode, this command causes reductions to stop when the reductants get to containing subterms that match the given term. If no term is given, this restriction is lifted.
Related: step switch
CafeOBJ> open NAT .
%NAT> set step on .
%NAT> set stop pattern s 2 .
%NAT> red s s s s s s s s s 0 .
>> target: (s 0)
STEP[1]? c
>> term matches to stop pattern: (s 2)
<< will stop rewriting
>> stop because matches stop pattern.
>> target: (s 2)
STEP[3]? c
(9):NzNat
Switches control various aspects of the computations and behavior of CafeOBJ. The current list of switches and their values can be shown with
show switches
The single switches are described separately in this manual.
:theory <op_name> : <arity> -> <coarity> { assoc | comm | id: <term> }
Adds operator theory 'associativity', 'commutativity', and/or 'identity' to an operator specfied by '
trace [whole]
switchDuring evaluation, it is sometimes desirable to see the rewrite sequences, not just the results. Setting the switch trace whole
will result in the resultant term of each rewrite step being printed. Setting the switch trace
will result in the display of which rule, substitution, and replacement are used.
trans [ <label-exp> ] <term> => <term> .
Defines a transition, which is like an equation but without symmetry.
See eq
for specification of requirements on <label-exp>
and the terms.
Transitions and equations server similar, but different purpose. In particular, reductions (both with or without behavior axioms used) do not take transitions into account. Only exec
also uses transitions. On the other hand, the built-in search predicate searches all possible transitions from a given term.
unprotect <module-name>
Remove overwrite protection from a module that has been protected with the protect
call. Some modules vital for the system are initially protected.
Related: protect
:use (<label> ... <label>)
Incorporate discharged goal sentences as new axioms.
using ( <modexp> )
Imports the object specified by modexp
into the current module without any restrictions on the models. See module expression
for format of modexp
.
Related: protecting
, including
, extending
var <var-name> : <sort-name>
Declares a variable <var-name>
to be of sort <sort-name>
. The scope of the variable is the current module. Redeclarations of variable names are not allowed. Several variable of the same sort can be declared at the same time using the vars
construct:
vars <var-name> ... : <sort-name>
Related: on-the-fly
, qualified term
, op
:verbose { on | off }
Turns on verbose reporting of the CITP subsystem.
Related: citp
verbose
switchPossible values: on
off
, default off
.
If turn on
, the system is much more verbose in many commands.
Related: quiet switch
version
Prints out the version of CafeOBJ.
view <name> from <modname> to <modname> { <viewelems> }
A view specifies ways to bind actual parameters to formal parameters (see parameterized module). The view has to specify the mapping of the sorts as well as the operators.
The <viewelems>
is a comma-separated list of expressions specifying these mappings:
sort <sortname> -> <sortname>
hsort <sortname> -> <sortname>
op <opname> -> <opname>
bop <opname> -> <opname>
and also can contain variable declarations.
Infix operators are represented as terms containing the operator with either literal underscores _
, or variables: _*_
or X * Y
. The <opname>
can be qualified.
In specifying views some rules can be omitted:
If the source and target modules have common submodules, all the sorts and modules declared therein are assumed to be mapped to themselves;
If the source and target modules have sorts and/or operators with identical names, they are mapped to their respective counterparts;
If the source module has a single sort and the target has a principal sort, the single sort is mapped to the principal sort.
Assume a module MONOID
with sort M
and ops e
and *
are given, and another SIMPLE-NAT
with sort Nat
and operators 0
and +
(with the same arity). Then the following expression constitutes a view:
view NAT-AS-MONOID from MONOID to SIMPLE-NAT {
sort M -> Nat,
op e -> 0,
op _*_ -> _+_
}