Higher-Order Rewriting with Types and Arities

Transcription

Higher-Order Rewriting with Types and Arities
Higher-Order Rewriting with Types and Arities?
Jean-Pierre Jouannaud1?? , Femke van Raamsdonk2 , and Albert Rubio3
1
LIX, École Polytechnique
91400 Palaiseau, France
email: [email protected] http: //www.lix.polytechnique.fr/˜jouannaud
2
Faculty of Sciences, Vrije Universiteit
De Boelelaan 1081a, 1081 HV Amsterdam, The Netherlands
email: [email protected]
3
Technical University of Catalonia
Pau Gargallo 5, 08028 Barcelona, Spain
email: [email protected]
1
Introduction
Background. Higher-order rewrite rules are increasingly used in programming
languages and logical systems, with three main goals: to set up a general infrastructure for computations, to describe rule-based decision procedures, and
to encode other logical systems. As usual, the choice of the syntax in which the
rules are expressed plays a crucial role, as well as the definition of rewriting itself.
Several proposals for frameworks of higher-order rewriting (we will not consider
here the explicit substitutions calculi) have been made by Klop, Khasidashvili,
Nipkow, and by Jouannaud and Okada. We will compare them with an example:
the encoding of λ-calculus with the β-reduction rule (λx. M )N → M {x 7→ N }.
Note that here the substitution {x 7→ N } is defined on a meta-level, that is, the
expression M {x 7→ N } is not a λ-term. In an encoding of logic, it is necessary
to internalize the definition of substitution.
The first proposal for a format of higher-order rewriting are the Combinatory
Reduction Systems (CRSs) defined by Klop [?], inspired by the Contraction
Schemes of Aczel [?]. The encoding of the β-reduction rule as a CRS is as follows:
app(abs([x]Z(x)), Z 0 ) → Z(Z 0 ). Here app and abs are function symbols, the
first of arity 2 and the second of arity 1. Both Z and Z 0 are so-called metavariables; they stand for the free variables to be instantiated when applying
the rule. Like function symbols they come equipped with an arity. Here Z has
arity 1 and Z 0 has arity 0. The abstraction operator [ ] is built-in. The rewrite
step (λx. yx)v → yv in the untyped λ-calculus is obtained using the following
substitution σ: {Z 7→ λu. app(y, u), Z 0 7→ v}. In general, a substitution τ assigns
to a meta-variable U of arity n a so-called substitute of the form λ(x1 , . . . , xn ). t.
The result (written in postfix notation) of applying τ to a term U (t1 , . . . , tn )
?
??
This work was partly supported by the RNTL project CALIFE and the CICYT
project HEMOSS, ref. TIC98-0949-C02-01.
Project LogiCal, Pôle Commun de Recherche en Informatique du plateau de Saclay,
CNRS, École Polytechnique, INRIA, Université Paris-Sud.
is t[x1 := t1 τ, . . . , xn := tn τ ], that is, it is the result of performing a complete
development in the λ-calculus with underlined λ’s. In the example, applying
the substitution σ yields the rewrite step app(abs([x]app(y, x)), v) → app(y, v).
Independently of Klop, Khasidashvili [?] introduced the framework of Expression
Reduction Systems (ERSs) which are similar to CRSs. Because it came first
on the market, the formalism suffers two drawbacks: a complex definition of
rewriting, where a specific meta-language is used, and the lack of types, making
it impossible to encode typed lambda-calculi.
The second proposal are the Higher-Order Rewrite Systems (HRSs) introduced by Nipkow [?,?]. The β-rule of the untyped λ-calculus is encoded in an
HRS as follows: app(abs(λx.zx), z 0 ) → zz 0 , juxtaposition denoting application
as usual. Here we work with simply typed λ-calculus as a meta-language. We
have app : term → term → term and abs : (term → term) → term with term
the type of the encoded (untyped) lambda terms. Further, z : term → term and
z : term are free variables, to be instantiated in order to get a rewrite step.
Pattern-matching is modulo βη, this is called higher-order pattern matching.
The encoding of the step (λx. yx)v → yv is obtained by using the following
substitution: {z 7→ λu. app(y, u), z 0 7→ v}. Since in HRSs we work with β-normal
forms (where β-reduction is the reduction in the meta-language), we get the following step: app(abs(λx. app(y, x)), v) → app(y, v). Working modulo βη makes
the approach quite powerful. As CRSs, HRSs suffer two weak points. First, rules
must be of base type. The reason is that by admitting rules of functional type,
one looses the so-called Church-Rosser property: the equality relation induced
by the rewrite rules viewed as equations together with βη no longer coincides
with the reflexive-transitive-symmetric closure of the rewrite relation. Note that
because in HRSs one works with long βη-normal forms, the left- and right-hand
side of a rule of functional type are of the form λx. s. For example, rules for
differentiation are disallowed. A way out is to have the less natural rule calculating the differential in a point: diff(λx. sin(x), u) → cos(u), a technique which
does not suffice to encode typed lambda-calculi. Second, rules must be written
in η-long form, which may be cumbersome, as for instance in the following rule
for map: map(λx. (F x), cons(h, t)) → cons((F h), map(λx. (F x), t) instead of the
more natural map(F, cons(h, t)) → cons(F (h), map(F, t)).
The third proposal is due to Jouannaud and Okada [?,?]. In AlgebraicFunctional Systems (AFSs) the rewrite relation is generated by a set of algebraic rules together with the β-reduction rule. The latter is written as follows:
{a : ∗, b : ∗, u : b, v : a, app : (a → b) × a → b} ` app(λx : a.u, v) → u{x 7→ v}.
Being polymorphic since the type of both sides is the type variable b, this rule
has instances of higher-order type. The goal here is to capture and extend the
kind of (typed) rewrite rules present in the calculus of inductive constructions
and automate the termination proof of such calculi by using specific recursive
schemas [?] or ordering techniques [?]. Weak points of this framework are the
use of plain pattern matching for firing rules, which is too weak for most logical
encodings, and the fact that the right-hand side of beta reduction is not a term.
2
Problem. Two questions are considered in this paper. How to allow rules of higher
type in Nipkow’s framework without compromising the Church-Rosser property
is the first. The answer originates from earlier work on rewriting modulo equationnal theories [?,?,?]. Since Nipkow’s framework still suffers using variables
without arity, the second question considered here is how to define typed higherorder rewriting modulo βη (with higher-order pattern-matching) where function
symbols and variables have types and arities such that higher-order rewrite rules
can be allowed. Some natural further questions are then how pattern unification
and termination methods can be adapted to this setting. A more fundamental
question, not addressed in the present paper, is whether the setting scales up to
richer type systems including dependent types.
Contribution. Our first main contribution is a careful investigation of normalform rewriting with a set of rules R modulo a convergent system S coming with
a matching algorithm and a complete unification algorithm. Normal rewriting is
defined as rewriting terms in S-normal form with rules in R by using S-pattern
matching. Under the minimal assumption that rewriting (with S and with R)
terminates , we show that the Church-Rosser property can be reduced to the
existence of extended rules “a la Stickel” [?] and the joinability of critical pairs
of various kinds: within rules in R using S-unification, and by sometimes also
overlapping rules in S with rules in R. Nipkow’s framework uses long-βη-normal
forms for S together with higher-order pattern matching and unification for
patterns. He does not need extensions nor checking critical pairs between S and
R because there are none under his assumption that rules are of basic type, an
assumption which can be removed when adding β-extensions, as we show.
Our second main contribution is a new framework for higher-order rewriting
that combines the advantages of the three earlier introduced frameworks. Rewriting is modulo βη but uses βη-normal forms instead of long-βη-normal forms as
representatives of a βη-equivalence class. For instance, let sin of type R → R be
a function symbol of arity 1 (written sin : R ⇒ R), and diff for differentiation
with type (R → R) → R → R be another function symbol of arity 1 (written
diff : (R → R) ⇒ R → R). Because of these arities, the terms λx. sin(x) and
diff(λx. sin(x)) are in βη-normal form. Function symbols are required to have as
many arguments as prescribed by their arity, not by the arity of their type, as is
the case of HRSs, where its long-η-normal form λu. @(diff(λx. sin(x)), u) should
be considered instead of diff(λx. sin(x)). Differentiation ca now be expressed with
rules whose lefthand side is always headed by a function symbol, as in
diff(λx. sin(x)) → λx. cos(x)
No extension rule is needed to ensure the Church-Rosser property of such rules.
We also tailor the definition of patterns due to Miller [?] to our case where
terms are not in long-βη-normal form, but in βη-normal form and respecting
the arity. A decision procedure for the unification of patterns is easily adapted
to this case.
A major advantage of our framework compared to Nipkow’s lies in the termination property needed: while there is no method known to prove termination
3
of R and S reductions together when S defines long-βη-normal forms, such a
method exists for the case where S uses both β and η as reductions, based on
the the higher-order recursive path ordering of Jouannaud and Rubio [?]. In a
companion paper [?], a new version of the higher-order recursive path ordering
is given which improves over the previous one.
For the sake of simplicity, we decided to develop our framework with a simple type system, omitting polymorphism as well as dependant types. Adding a
restricted form of polymorphism as in [?] is an easy step. Adding full polymorphism as well as dependent types is left for future work.
2
Terms and types
In this section, we introduce the basic notions and notations about (functional)
types, variables with type and arity, first- and higher-order terms, substitutions,
and first- and higher-order rewrite rules. We will in particular discuss the two
kinds of normal forms related to the use of the η-rule as a reduction or an
expansion which play an important role in the rest of the paper.
Types. We assume a set of sort symbols written as s, s0 , . . .. Every sort symbol
has a fixed arity. Further we assume a binary type constructor →. The set of
types is inductively defined by the following grammar:
σ, τ ::= s(σ1 , . . . , σn ) | σ → σ 0
where n is the arity of the sort symbol s. We write s instead of s() if the arity
of s is zero. For instance, we can represent the natural numbers by the type nat
with nat a sort symbol of arity zero. If list is a sort symbol of arity 1, it can be
used to build the type list(nat) of finite lists of natural numbers, and the type
list(nat → nat) of finite lists of functions on natural numbers. A type is said to
be functional if its head symbol is →, and a data type if its head symbol is a
sort symbol. A base type is a type without →.
Signature. We assume a set V of variables written as X, Y, Z, . . .. We distinguish
a set X ⊆ V of variables that can be abstracted over. For variables in X we
use besides the notation X, Y, Z, . . . also x, y, . . .. We assume a set F of function
symbols written as f, g, . . .; we sometimes also use more suggestive notations.
A signature consists of two parts: a set of variables V with a distinguished
subset X , and a set of function symbols F. For building terms, variables and
function symbols are almost treated in the same way. The main difference is that
(some) variables can be abstracted over while function symbols cannot.
Raw terms. The raw terms are built from a signature using application (denoted by @) and abstraction (denoted using λ). That is, the set of raw terms is
inductively defined by the following grammar:
s ::= X(s1 , . . . , sn ) | f (s1 , . . . , sn ) | λx : σ. s | @(s, s0 )
4
Note that the abstraction is over variables in X only. (Recall that the notation x
is only for variables in X .) A term of the form λx : σ. s is called an abstraction.
All occurrences of x in s are bound in λx : σ. s, and we work modulo renaming
of bound variables (α-conversion) as usual. Types occur in abstractions, but at
this point are completely arbitrary. We often omit the type in an abstraction,
writing λx. s instead of λx : σ. s. The term @(s, t) is said to be the application
of s to t. We write f for f () and X for X(). A term built solely from variables
and function symbols is said to be algebraic.
It is convenient to see terms as trees in which a subterm has a precise position
denoted by a sequence of natural numbers. In this view, each abstraction λx : σ
is seen as a particular symbol of arity one. We will use the notation s|p for the
subterm at position p in s, and s[u]p for the term obtained from s by replacing
s|p by u (variables can be captured here).
Type declarations. In our framework, function symbols and variables have besides a type also an arity which is a natural number denoting the number of
arguments it is supposed to have. That is, we combine a typed setting as in
HRSs and AFSs with the presence of arities as in CRSs (all variables have arity
0 in AFSs and HRSs). In order to express the combination of a type and an arity,
we use type declarations as in AFSs. Type declarations are not types, but they
are used for typing purposes. A type declaration is an expression of the form
σ1 × . . . × σn ⇒ τ
with σ1 , . . . , σn , τ types. Such a type declaration is a notation for the type σ1 →
. . . → σn → τ in which it is also expressed that the arity is n. The types
σ1 , . . . , σn are said to be the input types, and τ is said to be the output type. For
a type declaration ⇒ τ of arity zero we also sometimes write τ . Note that the
output type τ may be functional. We use υ, υ 0 , . . . to denote type declarations. We
also use σ, τ, . . . to denote something which is either a type or a type declaration.
A type declaration is first-order if it doesn’t use functional types, and higherorder if it does.
Environments. An environment is a finite set of pairs written as {q1 : υ1 , . . . , qn :
υn } with each qi a variable or a function symbol, and each υi a type declaration. All qi are assumed to be distinct variables or function symbols. The set
{q1 , . . . , qn } is the domain of the environment Γ . Environments are denoted by
Γ, ∆, . . ..
Typing rules. We use the following typing rules to derive judgements of the form
Γ ` t : σ, with Γ an environment, t a term, and σ a type:
X : σ1 × . . . × σn ⇒ τ ∈ Γ
Γ ` si : σi
Γ, x : σ ` t : τ
var
abs
Γ ` X(s1 , . . . , sn ) : τ
Γ ` λx : σ. t : σ → τ
f : σ1 × . . . × σn ⇒ τ ∈ Γ
Γ ` si : σi
Γ ` t : σ app
f un Γ ` s : σ → τ
Γ ` f (s1 , . . . , sn ) : τ
Γ ` @(s, t) : τ
5
Concerning the var rule: note that X ∈ V and X ⊆ V so X may or may not be
in X . If X ∈ X , then n = 0. Concerning the var and fun rules: the typing system
guarantees that function symbols and variables always get the exact number of
arguments as prescribed by their arity. Therefore we use in the remainder of the
paper the following convention:
The notations f (s1 , . . . , sm ) and X(s1 , . . . , sm ) imply that f and X have
arity m. Likewise, writing @(f, a) means that f has arity 0 with a type
declaration of the form f :⇒ τ1 → τ2 .
Concerning the abstraction rule: recall that we can only abstract over variables of arity 0; indeed by the notation convention x ∈ X . Bound variables in
our framework correspond to the bound variables in HRSs, and to the object
variables (that may or may not be bound) in CRSs. On the other hand, the free
variables in the rewrite rules may have an arbitrary arity. They correspond to
the free variables in HRSs, and to the meta-variables in CRSs. We feel that having no meta-variables is important both conceptually and practically, if rewrite
rules may themselves become objects of a computation.
Term substitutions. Given an environment Γ , a term substitution γ, or simply
substitution, is a mapping from a finite subset of the variables declared in Γ
to terms, written as a finite set of pairs {X1 7→ s1 , . . . , Xn 7→ sn }, satisfying a
type-preservation property:
∀Y 7→ t ∈ γ with Y : τ1 × . . . × τk ⇒ τ ∈ Γ , then si = λy1 : τ1 . . . . .λyk : τk . u,
with Γ ∪ {y1 : τ1 , . . . , yk : τk } ` u : τ , assuming that every variable yi is taken
fresh from X .
We use postfix notation for substitution application, and therefore Y γ has
the substitute t for value. A term substitution is extended to an endomorphism
from terms to terms, as follows:
X(s1 , . . . , sn )γ
f (s1 , . . . , sn )γ
(λx : σ. s)γ
@(s, s0 )γ
= @(Xγ, s1 γ, . . . , sn γ)
= f (s1 γ, . . . , sn γ)
= λx : σ. sγ
= @(sγ, s0 γ)
In the clause for abstraction we assume that the name of the bound variable is
chosen such that it is disjoint from the domain and range of the substitution
γ. We see here that X(s1 , . . . , sm ) in fact denotes the application of X to its
arguments. If A = {s1 , . . . , sm } is a set or a multiset of terms, then we also use
the notation Aγ to denote the set or multiset {s1 γ, . . . , sm γ}.
Proposition 1.
If Γ ` s : σ then Γ ` sγ : σ.
Proof. The proof proceeds by induction on Γ ` s : σ. For the case of the variable
rule, we use our assumption that substitutions preserve types.
6
Modulo βη. As already remarked above, we work modulo renaming of bound
variables, or α-conversion. In addition, we work modulo two specific equations:
@(λx : a. s, t) =β s{x 7→ t}
λx : a. @(s, x) =η s
as in the HRSs defined by Nipkow. As his, our framework uses higher-order
pattern matching. As pointed out by Nipkow, there is no need to explicit the
side condition that x does not occut free in s if s is considered as a free variable,
since it can only be substituted by terms not containing the bound variable x
by the definition of substitution in an abstraction.
Up to now, Nipkow’s framework is simply a particular case of ours, obtained
by restricting the set of variables to X , that is, all variables have arity 0. The
importance of variables of non-zero arity comes next.
In HRSs one uses as the representative of the βη-equivalence class of a term
t its (unique) long-βη-normal form denoted by t lηβ , obtained by orienting the
β-equation from left to right, yielding the β-reduction rule @(λx : a. s, t) →β
s{x 7→ t}, and the η-equation from right to left, yielding the η-expansion rule
s →η−1 λx : a. @(s, x). Of course, the use of η-expansion must be restricted
to ensure termination. Together, β-reduction and restricted η-expansion form a
convergent rewrite system, whose normal forms are precisely the long-βη-normal
forms.
In contrast, we use as unique representative of the βη-equivalence class of a
term t its βη-normal form denoted by t↓βη (or simply t↓), obtained by using
both equations from left to right, replacing therefore the η-expansion rule by
the η-reduction rule: λx : a. @(s, x) →η s. Again, It is well-known that β- and
η-reductions form a convergent rewrite system, whose normal forms are precisely
the βη-normal forms.
The interaction between algebraic symbols and the η-reduction rule is worth
an explanation, since it is very different from the case of HRSs. Because of the
typing system, the symbol f in the term λx. @(f, x) has arity 0. Here we can
apply the η-reduction rule which yields the term f . In contrast, if f is a symbol
of arity 1, then λx. @(f, x) is not a term: the corresponding term is λx. f (x),
but here the η-reduction rule doesn’t apply. We see that our representatives are
actually in between usual η-normal and η-expanded forms.
Patterns. Unification of simply typed λ-terms is undecidable, but Miller [?]
defined a class of simply typed λ-terms, the patterns, for which unification is
decidable and a most general unifier can be computed in linear time.
We consider the natural adaptation of the definition of pattern for our framework, as an application-free term S in βη-normal form such that for every subterm of the form X(y1 , . . . , yn ) where X is a free variable of arity n, the yi are
different variables bound in s.
It is easy to see that patterns can be defined via a typing system in which
judgements are of the form ∆; Γ ` s : σ as in [?]. The intuition is that ∆ contains
the type declarations for the free variables and the function symbols, and that
Γ contains the declarations for the variables from X that are still free in s, but
7
should be bound in the complete term where s occurs as a subterm. A term p is
a pattern if we can derive ∆; ∅ ` p : σ. This definition of patterns is convenient
for unification.
3
Rewriting modulo
A main problem with Nipkow’s HRSs is that rules of higher type are ruled
out. This is so because the absence of critical pairs, though ensuring confluence,
does not ensure the Church-Rosser property. These problems actually originate
in the notion of rewriting with a set R of rules modulo a set S of equations
built in the pattern-matching algorithm. There are four approaches to rewriting
modulo, corresponding to four possible definitions of rewriting in this context.
All have been developped for first-order rewriting, but the methods and results
are essentially general.
In the rest of this section, we assume given two sets: a set R of rules and a set
S of equations. In pratice, S may contain associativity and commutativity in the
case of first-order computations, and βη in the case of higher-order computations.
Recall that (plain) rewriting is defined as
p
s −→ t
R
3.1
if s|p = lσ and t = s[rσ]p for some rule l → r ∈ R
Class-rewriting
This approach is due to Lankford [], who investigated rewriting modulo permutative theories like commutativity. The rewrite relation is defined as
s −→ t
RS
if s =S u −→ t
R
The Church-Rosser properties are defined as follows:
Definition 1. A pair (s, t) of terms is said to be joinable with class rewriting
if there exists two terms u, v such that s −→∗RS v =S w ←−∗RS t. Class rewriting
is
(i) Church-Rosser if ∀s, t such that s ←→∗R∪S t, the pair (s, t) is joinable.
(ii) locally-confluent if ∀s, t, u such that u −→RS s and u −→RS t the pair
(s, t) is joinable.
Under the termination assumption of class rewriting, the Church-Rosser
property is equivalent to local confluence, an easy exercice based on Newman’s
lemma. This approach, unfortunately, implies searching the equivalence class of
the term to be rewritten, which must therefore be finite. However, even when
equivalence classes are finite, the search can be quite time consuming. There is
indeed a general agreement that this solution is not feasible, although simple to
carry out.
8
3.2
Plain-rewriting modulo
This approach is due to Huet, who tried to remedy to the weaknesses of Lankford’s approach by using plain rewriting.
Definition 2. A pair (s, t) of terms is said to be joinable with plain-rewriting
modulo if there exists two terms u, v such that s −→∗R v =S w ←−∗R t. Plain
rewriting modulo is
(i) Church-Rosser if ∀s, t such that s ←→∗R∪S t, the pair (s, t) is joinable;
(ii) locally-confluent if ∀s, t, u such that u −→R s and u −→R t, the pair (s, t)
is joinable;
(iii) locally-coherent if ∀s, t, u such that u −→R s and u ←→S t, the pair (s, t)
is joinable.
Under the assumption that class-rewriting terminates, the Church-Rosser
property is equivalent to local-confluence and local-coherence together. When it
comes to critical pairs, local-confluence and local-coherence can be reduced to
the joinability of critical (confluence and coherence) pairs. This approach suffers
a severe restriction: rewrite rules in R must be left-linear, while equations in S
must be left- and right-linear. Note that −→R ⊆ −→RS .
3.3
Rewriting modulo
This approach is due to Paterson and Stickel [] who introduced it in the context
of AC-rewriting. The theory of rewriting modulo was later done by Jouannaud
and Kirchner who treated the general case []. The idea is to make the search of
a rewritable term efficient while eliminating the linearity restrictions, thanks to
rewriting modulo defined as:
p
s −→ t
RS
if s|p =S u|p
p
such that u −→ t
R
Here, the use of the equations in S is allowed below (and only below) the position
p at which the rewrite step takes place. As a consequence, and letting l → r be
the rule used in the rewrite step, we have s|p =S u|p = lσ, and therefore s|p =S lσ
meaning that σ can be computed by matching modulo S. It turns out that Smatching is decidable for permutative axioms, as well as for associativity and
commutativity together. Similarly, βη-matching, called higher-order matching,
is known to be decidable when the terms to be unified are patterns [?], in which
case it is even of linear complexity, while its decidability is still open in the
general case. Note again that −→R ⊆ −→RS .
Definition 3. A pair (s, t) of terms is said to be joinable with rewriting modulo
if there exist two terms u, v such that s −→∗RS v =S w ←−∗RS t. Rewriting modulo
is
(i) Church-Rosser if ∀s, t such that s ←→∗R∪S t, the pair (s, t) is joinable;
(ii) locally-confluent if ∀s, t, u such that u −→RS s and u −→R t, the pair (s, t)
is joinable;
(iii) locally coherent if ∀s, t, u such that u −→RS s and u ←→S t, the pair
(s, t) is joinable.
9
Under the assumption that class-rewriting terminates, the Church-Rosser
property is equivalent to local-confluence and local-coherence together. Note the
use of both −→RS and −→R in the definition of local-confluence. When it comes
to critical pairs, local confluence can be reduced to the joinability of complete
sets of critical pairs, defined by using S-unification instead of plain unification:
Definition 4. Given two rules l → r and g → d such that l|p unifies with g with
a complete set of S-unifiers Σ, the set {lσ[dσ]p , rσ | σ ∈ Σ} is called a complete
set of S-critical pairs of g → d on l → r at position p. We call CP (R, S) the set
of all S-critical pairs, for all rules in R.
Local confluence reduces to the joinability of all pairs in CP (R, S). When unification is finite, as it is the case for AC, or for higher-order unification of patterns,
local confluence is therefore decidable under our assumptions.
Local-coherence can also be checked by computing adequate critical pairs,
but there is a cheaper way out based on unifiability. Let l → r be a rule in R
and g = d be an equation in S such that g|p and l unify modulo S at a position
p in g which is neither the root position, nor a variable position, then the rule
g[l]p → g[r]p is called an S-extension of l → r. Now, assuming σ is an S-unifier
of g|p and l, we have:
Λ
(≥p)∗
p
dσ ←→ gσ = gσ[g|p σ]p ←→ gσ[lσ]p −→ gσ[rσ]p
S
R
and it is easy to see that dσ −→Λ
6 pg→r gσ[rσ]p
g[l]→g[r] gσ[rσ]p , but that dσ −→
0
since there is an equational step above p. Note that any rule l → d0 such that
l0 ←→∗S g[l] and d0 ←→∗S g[r] would do as well. In case S is a terminating rewrite
system, it will indeed be appropriate to S-normalize g[l] and g[r].
Assuming that R is closed under the computation of extensions, local-coherence
therefore holds [?]. Besides, extensions of extensions are not needed in the ACcase [?].
3.4
Normalized rewriting
The previous approach has two major drawbacks: the hypothesis that classrewriting terminates while rewriting modulo is actually used; a related additional
requirement that congruence classes are finite. To remedy this situation, a last
approach was promoted by Claude Marché [], which appplies when the equations
can be oriented into a terminating Church-Rosser (possibly modulo AC) set of
rules (again denoted by S), so as to get normal forms (possibly modulo AC).
The point here is that a weaker termination property is needed: termination of
AC-class rewriting with R∪S, instead of termination of (S ∪AC)-class rewriting
with R. Besides, the finiteness requirement for S-congruence classes vanishes.
Marché’s normalized rewriting is defined on arbitrary terms as s −→pRS↓ t
iff s −→!SAC s ↓S −→RAC t. Of course, in case of absence of the AC axioms,
plain pattern-matching is used. Marché shows that the confluence property of
10
normalized rewriting can be reduced to critical pairs as previously. These critical
pairs are computed with a complete unification algorithm for a theory T which
can be chosen arbitrarily between AC and S ∪ AC. In pratice, the choice of T
depends on the behaviour of the unification algorithm: it is well known that AC
has exponentially more general unifiers than ACZ, where Z denotes the identity.
It is therefore useful to incoporate Z to the unification algorithm when Z ∈ S.
There are also, of course, AC-extensions. Besides, there are what Marché calls
S-normalizing pairs, which come from the possible superpositions of rules in S
within rules in R which make instances of rewrite rules in R reducible by rules
in S. We skip the complex definition of these pairs here.
Because AC-matching is used instead of S∪AC matching, Marché’s approach
is indeed a generalization of Huet’s approach rather than one of Peterson and
Stickel’s. When it comes to higher-order rewriting, it becomes crucial to use
higher-order pattern matching (or AC-higher-order pattern matching [?]) for
firing rules instead of plain pattern matching (or AC-pattern matching). This is
not within the scope of Marché’s results.
Finally, a rather different approach appears in [?], in which the rewrite relation −→∗S −→R is used, with the argument that normalized rewriting is a particular strategy of the former. This is a fact, but the argument does not really
hold since it leads to much weaker characterizations of confluence.
3.5
Higher-order rewriting
None of the previous appproaches apply to higher-order rewriting for the following reasons:
– Class-rewriting does not apply because βη-equivalence classes are infinite;
– Plain-rewriting modulo does not apply because critical pairs between the
right-hand side of the β-rule and rules in R are not computable since the
β-rule is a schema with infinitely many instances;
– Rewriting modulo does not apply because extended pairs with the righthand
side of the β-rule are not computable for the same reason as previously;
– Normal rewriting does not apply because it is not based on S-matching for
firing rules.
As a conclusion, Nipkow’s results cannot be seen as an instance of any of
these approaches.
4
Church-Rosser properties of normal rewriting
In this section, we give an original abstract presentation of rewriting “‘a la
Nipkow”, which we will call normal rewriting to distinguih it from Marché’s
normalized rewriting. We then review and generalize Nipkow’s results before to
study our own framework for higher-order rewriting. Applying our results to
rewriting modulo identity, or identity and idempotency, or commutative group
11
theory (in order to compare our results with Marché’s in these cases), or higherorder AC-rewriting (to cover this practically important case) would require a
generalization of these results to the case where S is a convergent set of rules
modulo AC. We believe that this should not be too difficult but we have not
carried out the details yet.
4.1
Abstract results
We assume that S is a set of rules, and define normal rewriting as a relation on
terms in S-normal form as follows:
p
s −→ t
RS↓
if s = s↓S , s|p =S lσ for some l → r ∈ R, and t = (s[rσ]p )↓S
We could of course define normal rewriting on arbitrary terms as s −→pRS↓ t iff
s −→!S s↓S −→RS t, to the price of cosmetic differences.
General assumptions for normal rewriting
In the rest of this section, we assume the following properties of the sets R, S:
(a) S is Church-Rosser set of rules,
(b) RS ∪ S is terminating,
(c) rules in R are S-normalized.
Assumptions (a) and (b) together imply that any term has a unique Snormal form while assumption (c) comes for free. Assumption (c) is essential for
the applications. We now come to the Church-Rosser properties:
Definition 5. A pair (s, t) of terms is said to be joinable with normal rewriting
if there exists a term v in normal form such that s↓S −→∗RS↓ v ←−∗RS↓ t↓S (resp. s↓S
−→∗RS↓ t↓S ).
Normal rewriting is
(i) Church-Rosser if ∀s, t such that s ←→∗R∪S t, the pair (s, t) is joinable;
(ii) locally confluent if ∀t, u, v with t in S-normal form such that t −→RS u
and t −→RS v, the pair (u, v) is joinable;
(iii) locally coherent if ∀s, t such that s −→RS t, the pair (s, t) is joinable,
Note that these Church-Rosser properties use rewriting or rewriting modulo in their hypotheses, and normal-rewriting in their conclusions. Note also
that local confluence assumes that the term s is in S-normal form. Thanks to
assumptions (a,b), this allows us to reformulate local confluence as follows:
Lemma 1. Normal rewriting with (R, S) is locally confluent iff ∀s, t, u, v such
(≥p)∗
that s −→pR u, s −→S
t and t −→qRS v with p#q or q ≥ p and t in S-normal
form, then the pair (v, u) is joinable.
Proof. Assume that t −→pRS u and t −→qRS v. In case p, q are comparable in the
prefix ordering of strings, we choose q ≥ p. Then, since t −→pRS u, there exists a
(≥p)∗
term s such that s −→pR u and s ←→S
t. Using assumption (a) and the fact
that t is in S-normal form yields the conclusion.
2
12
Theorem 1. Assuming local confluence and local coherence, normal rewriting
is Church-Rosser.
Proof. Let s, t be two terms such that s ←→∗RS ∪S t, an assumption equivalent
to the Church-Rosser assumption s ←→∗R∪S t, but more convenient for our induction argument. We show that the pair (s, t) is joinable by induction on
the multiset of terms occuring in the proof between s and t ordered by ≥mul
with ≥= (−→RS ∪ −→S )∗ . We will freely use the fact that −→RS↓ ∪ −→S ⊆
(−→RS ∪ −→S )+ .
The case where s = t being trivial, let us consider the case s ←→∗R∪S u ←→R∪S t.
By induction hypothesis, the pair (s, u) is joinable, hence s ↓S −→∗RS↓ v and
u↓S −→∗RS↓ v for some normal form term v. We distinguish cases according first
to the last step between u and t and then to the S-reducibility of u:
1. Let t ←→S u. By assumptions (a,b), t↓S = u↓S and we are done.
2. Assume that u ←→RS t and u 6= u↓. By local coherence assumption, the pair
(u, t) is joinable, that is u↓S −→∗RS↓ w and t↓S −→∗RS↓ w for some w. The
proof v ←−∗RS↓ u↓S −→∗RS↓ w can be seen as a proof with S-steps and RS
steps made of terms which are all strictly smaller than s. We can therefore
conclude by induction hypothesis.
3. Let t −→RS u and u = u↓S . By local coherence assumption, the pair (u, t) is
joinable, that is u −→∗RS↓ w and t↓S −→∗RS↓ w for some w. Now, the proof
v ←−∗RS↓ u −→∗RS↓ w is made of terms all strictly smaller than t, hence is
smaller than the original proof. We conclude by induction hypothesis.
4. Assume that u −→RS t and u = u↓. If u = v, we are done since u −→RS↓ t↓S .
Otherwise, u −→RS v 0 −→∗S v 0↓S −→∗RS v. Since u is in normal form, by local
confluence assumption there exists a term w such that v 0↓S −→∗RS↓ w ←−∗RS↓ t↓S .
We now consider the proof v ←−∗RS↓ v 0↓S −→∗RS↓ w, which is made of terms
all strictly smaller than u. By induction hypothesis, v −→∗RS↓ ←−∗RS↓ w and
we are done.
2
Note that termination of −→S and −→RS↓ suffice for this proof.
We are left to show that local coherence and local confluence can be reduced
respectively to the use of extended rules and to the checking of critical pairs respectively. Since these pairs are quite similar to those used by rewriting modulo,
it becomes extremely easy to incorporate normal rewriting in existing rewriting
libraries.
Proposition 2 (Local Coherence). Let s −→RS t and let us assume that
(i) R is closed under the computation of normalized S-extensions,
(ii) the Church-Rosser property holds for proofs strictly smaller than {s, t}.
Then the pair (s, t) is joinable.
Note that assumption (ii) is satisfied when local coherence is used in the
proof of Theorem 1.
13
Proof. Assume first that s is in normal form with respect to S, then s ↓S =
s −→pRS t, and therefore s↓S −→pRS↓ t↓S and we are done.
Otherwise, s −→qg→d∈S u for some u. We distinguish several cases according
to the respective positions of p and q.
1. If p and q are disjoint, then t −→qS v and u −→pRS w for some w obtained by
reducing s at p and q with respectively RS and S. We conclude by assumption
(ii).
2. If q ≥ p, then u −→pRS t again by definition of rewriting modulo, and we
conclude by assumption (ii) again.
3. If p = q · p0 and p0 6∈ FDom(g), we have the following usual commutation
diagram: u −→∗RS v ←−qS ←−∗RS t. We conclude by assumption (ii) that the
pair (u↓, t) is joinable, and therefore the pair (s, t) is joinable as well.
4. Finally, if p = q·p0 and p0 ∈ FDom(g)\{Λ}, then u −→qRS t with the extended
rule g[l]p0 → g[r]p0 , and we conclude once more by assumption (ii).
2
In case S is a set of linear rules (both right- and left-linear), the proof can
be carried out by a simple induction on s ordered with −→S , and does not need
assumption (ii) anymore. Assuming only left-linearity, the result then needs a
stronger induction argument. In both cases, the conclusion is that s↓ −→∗RS↓ t↓.
Note that the possible degenerated case s↓= t↓, which may happen in case the
RS -redex disappears during the S-normalisation.
We now move to local confluence. A critical pair (rσ, lσ[dσ]p ) ∈ CP (R, S) is
said to be reducible if lσ is S-reducible.
The proof of local confluence will be based on the use of Lemma 1 which
provides an alternative definition of local confluence. The proof will of course be
done by induction on the number of S-steps squeezed between the R-step and
the RS -step. This number can be reduced either from the left or the right. This
suggests splitting the set S into two subsets, leading to two different notions of
critical pairs:
Definition 6. We say that S1 ∪ S2 = S is a splitting of S, iff for every term t,
there exists a derivation of the form t −→∗S1 t↓S1 −→∗S2 t↓S .
In lambda calculus, splitting corresponds to the postponment of the η-rule.
In section ??, we will put η in S1 or S2 depending on its orientation.
Definition 7. Given a rule g → d ∈ S2 , a rule l → r ∈ R, and a position
p ∈ FDom(d) \ {Λ} such that l S-unifies with d|p , then the rule d[l]p → d[r]p is
called a forward extension of R with S2 .
In particular, rules of the form g → x with x ∈ X or g → f (x) with x ∈ X ∗
have no forward extension. In practice, it is desirable that rules in S2 satisfy this
restriction.
As for extensions, it is easy to see that forward extensions satisfy their purpose: gσ −→Λ
RS dσ[rσ]p if σ is a unifier of the equation l = d|p .
This definition of forward extensions works fine with first-order rewriting,
but not with higher-order rewriting. The problem here is that the properties of
14
the higher-order structure are defined by rule schemas, like the β-rule. Taking
the infinitely many instances of the rule schema yields infinitely many forward
extensions. Besides, the righthand side of β being of the form u{x 7→ v}, many
different lefthand sides correspond to the same righthand side term. However,
we can give a more general formulation of the forwward extension of a rule, by
exploiting the expressive power of higher-order terms:
Definition 8. Given a rule g → d ∈ S2 , a rule l → r ∈ R, and a position
p ∈ FDom(d) \ {Λ} such that the equation X(l) = d has solutions modulo S,
then the higher-order rule X(l) → X(r) is a (higher-order) forward S2 -extension
of R.
Again, it is easy to see that higher-order forward extensions satisfy their
purpose if the β-rule belongs to S:
– Let us assume first as before that σ is a unifier of the equation l = d|p .
consider the substitution σ = {X 7→ λx.d[x]p }. Then, the β-normalized σinstance of the extension is the rule d[l]p → d[r]p as before, and will therefore
serve the same purpose.
– Let us now consider in the higher-order case, the case of the β-rule itself.
Assume that @(λx.u, v) −→β u{x 7→ v} −→p
w. The critical case
lτ →rτ mod S
is obtained when p ∈ FPos(u), but u|p is not an S-instance of l. Take σ =
{X 7→ λx.u{x 7→ v}[x]p } ∪ τ . Then, Xσ(lσ) −→β u{x 7→ v}[lτ ]p =S u{x 7→
v} ←−β @(λx.u, v), hence Xσ(lσ) =S @(λx.u, v) since the β-rule belongs
to S. Therefore @(λx.u, v) −→X(l)→X(r) mod S Xσ(rσ) −→β u{x 7→ v}[rτ ]p ,
which is again the expected result.
So, one higher-order forward extension covers all possible forward extensions
of the rule l → r. Besides, it does not need S-normalization (except for very
specific S that we will not encounter in practice), since l and r are already
S-normalized.
Note that the need for higher-order extensions arises with righthand sides of
rules, not with lefthand sides (this is why context variables were not needed for
the other extensions). This is because of the form of β wich takes two terms at
disjoint positions and stacks them up on the right. If we had a rule in S splitting
the lefthand side term into two parts written at disjoint positions on the right,
we would also need such extensions in this case. We have therefore implicitely
used the “geometrical” properties of the λ-calculus, as axiomatized in [8].
Definition 9. Given a rule l → r ∈ R, a rule g → d ∈ S1 such that g is not
a variable, and a position p ∈ FPos(l) such that l|p and g have a most general
plain unifier σ, then the pair (rσ, lσ[dσ]p ) is called a shallow critical pair of
g → d onto l → r at position p. We denote their set by SCP (S1 , R).
We say that a shallow pair (a, b) is strongly joinable if b −→∗S −→Λ
R c and the
pair (a, c) is joinable.
(≥p)∗
Proposition 3 (Local Confluence). Let s −→pl→r∈R u, s −→S
t and t −→qRS v
with the rule g → d ∈ R, p#q or q ≥ p, and t in S-normal form. Assuming that
15
(i) S1 -irreducible critical pairs in CP (R, S) are joinable
(ii) R is closed under the computation of normalized S-extensions,
(iii) S1 -irreducible shallow critical pairs in SCP (S1 , R) are strongly joinable,
(iv) R is closed under the computation of forward extensions with S2 ,
(v) the Church-Rossser property holds for proofs made of terms strictly smaller
than t,
then the pair (u, v) is joinable.
Note the form of assumption (v) using t (named u in the proof of Theorem 1
case 4). We cannot use s, of course, which does not show up in that proof, nor u
or v, since we cannot know which one plays the role of t (the other playing the
role of u↓S ).
Proof. The proof is by induction on the pair (s, n) where n is the length of the
S-derivation from s to t. Pairs are compared lexicographically, s is compared
with S-rewriting, and n with the ordering on natural numbers.
Assume that Var(l) ∩ Var(g) = ∅. Then, s|p = lσ, t|q =S gσ, u = s[rσ]p and
v =S t[dσ]q for some substitution σ.
The proof is in three parts. First, we reduce the number of S-steps between
s and t from the left. During this process, t remains the same, hence is kept in
S-normal form, while s is changed and becomes in S1 -normal form at the end
of the process. Second, we reduce the number of S-steps between s and t from
the right. In this process, t is changed, hence we loose the property that it is
in S-normal form, while s remains the same, hence is kept in S1 -normal form.
Finally, we consider the case where s = t, which is the base case of our induction.
In this case, s = t is therefore in S1 -normal form (hence in S-normal form in
case S2 = ∅).
0
∗
1. First part, in case s −→o≥p
a→b∈S1 s −→S t, for which there is at least one S1 step from s. Since o ≥ p by assumption, let o = p · o0 .
(a) o0 6∈ FPos(l). Then, the proof steps commute as follows: u −→∗a→b u0
(≥p)∗
and s0 −→a→b w −→pu0 for some terms u0 and w. By assumptions (a,b),
∗
(≥p)
w −→S
t since t is in S-normal form. By induction hypothesis, the
pair (u0 , v) is therefore joinable, and hence, so is the pair (u, v).
(b) o0 ∈ FPos(l). Since l is S1 -normalized, then a cannot be a variable.
Therefore, (u, s0 ) originates from a shallow critical pair which is strongly
joinable by assumption (iii). As a consequence, there exist terms w, w0
(≥p)∗
such that s0 −→S
w −→pRS w0 and the pair (u, w0 ) is joinable, that is
u↓S −→∗RS↓ u0 ←−∗RS↓ w0↓S . Since s −→+
S w, by induction hypothesis, the
0
0
pair (w , v) is joinable, that is w ↓S −→∗RS↓ v 0 ←−∗RS↓ v↓S .
Note that t −→pRS w0 , since all S-steps between t and w are at positions
≥ p. This shows that the proof u0 ←−∗RS↓ w0↓S −→RS↓ v 0 is made of terms
all strictly smaller than t. We can therefore close the diagram thanks to
assumption (v).
16
2. Second part, in case s −→∗S2 t0 −→oa→b∈S2 t, for which there is at least one
proof step left with S2 ending in t. We discuss according to the respective
positions of o, q.
(a) Case o#q. Then, t0 −→qRS v 0 −→oa→b v. By induction hypothesis, the pair
(u, v 0 ) is joinable, and by assumptions (a,b), the pair (u, v) is therefore
joinable.
(b) Case o ≥ q. Then t0 −→qRS v and we conclude by induction hypothesis.
(c) Case q = o · q 0 and q 0 6∈ FPos(g). Since Var(b) ⊆ Var(a), there exist
terms v, v 0 and w such that t0 −→RS v 0 −→∗RS −→oS w and v −→∗RS w.
By induction hypothesis, the pair (u, v 0 ) is joinable. We now conclude
by assumption (v).
(d) Case q = o · p0 and p0 ∈ FPos(g) \ {Λ}. Then, t|q = b|q0 σ =S gσ. By
assumption (iv), t0 −→ob[g]→b[d] v or t0 −→oX(g)]→X(d) v. We conclude by
induction hypothesis.
3. case s = t, for which t is in S-normal form in case S2 = ∅. We discuss as in
any critical pair lemma according to the respective positions of p and q.
(a) p and q are disjoint. Then both steps commute: u −→qRS w and v −→pR w
with w = s[rσ]p [dσ]q . We conclude by assumption (v) applied to the
proof u −→qRS w ←−pR v.
(b) q = p·q 0 and q 0 6∈ FPos(l). Similarly u −→∗g→dS w and v −→∗g→dS w0 −→pl→r w.
Note the absence of S-matching with the rule l → r: all S-instances of g
below l must be identical. We conclude by assumption (v) as previously.
(c) q = p·q 0 and q 0 ∈ FPos(l). Then, (s|p )|q0 = (l|q0 )σ =S gσ, and therefore,
the pair (l|q0 , g) is unifiable with some most-general unifier θ such that
σ =S θρ. Since s is in S-normal form, so is s|p . By assumption (i), the
critical pair (lθ[dθ]q0 , rθ) is therefore joinable. As a consquence,
u = u0 =S s[(lθ[dθ]q0 )ρ]p −→∗RS ∪S ←−∗RS ∪S s[(rθ)ρ]q = v 0 =S v. Since all
terms in the proof from u0 to v 0 are strictly smaller than s, we conclude
by assumption (v) and (a,b).
Note that our induction argument actually uses the termination of RS alone.
Putting these results together, we obtain:
Theorem 2. Let two sets R, S of rules satisfying properties (a), (b), (c), and
(S1 , S2 ) be a splitting of S. Assuming that
(i) R is closed under the computation of normalized extensions,
(ii) R is closed under the computation of forward pairs with S2 ,
(ii) S1 -irreducible critical pairs in CP (R, S) are joinable,
(iii) S1 -irreducible shallow critical pairs in SCP (R, S) are strongly joinable,
then normal rewriting is Church-Rosser.
This concludes our study of the Church-Rosser properties of normal rewriting. Before to apply our last result to various cases, we come back on our termination assumption.
17
4.2
Termination assumption
Going back to our definition of normal rewriting, a consequence of the fact that
normal-rewriting a term in S-normal form is nothing but rewriting it modulo S
followed by S-normalization is that termination of −→RS ∪ −→S implies termination of normal rewriting, and even of the relation −→RS↓ ∪ −→S . The converse
does not hold in general, but does hold under some specific assumption. First
we prove a simple lemma:
Lemma 2. Assume that S terminates, rules in S are left-linear, and there
are no overlaps of rules in S inside rules of R. Then, s −→pRS t implies s ↓S
−→∗RS↓ t↓S .
Proof. By induction on s with −→S . If s = s↓, the result follows by normalizing
t. Otherwise, let s −→ql→r∈S u. The are several cases depending upon the relative
positions of p and q.
1. If p#q, then u −→RS v and t −→S v for some v. e conclude by induction
hypothesis.
2. If q ≥ p, then u −→RS t, and we conclude by induction hypothesis again.
3. If p > q, then there is no overlap by our assumption. Therefore, by leftlinearity assumption, theer exists a term v such that u −→∗RS v and t −→S v.
We cvan now apply our induction hypothesis to the first step originating in
u, but not to the others!
Lemma 3. Assume that S is a Church-Rosser and terminating left-linear rewrite
system such that there are no overlaps of rules in S inside rules of R. Then normal rewriting −→pRS↓ terminates iff −→RS ∪ −→S terminates.
Proof. We are left showing the only if direction under our assumptions. Assume
an infinite derivation with −→RS ∪ −→S . Since S terminates by assumption,
there must be an infinite derivation of the form . . . un −→RS vn −→∗S un+1 −→RS . . ..
By the previous lema, un ↓ −→∗RS↓ vn ↓, and by the Church-Rosser assumption,
vn ↓= un+1 ↓, which generates an infinite derivation for −→RS↓ , unless all by
finitely many such steps are erased.
To conclude, we need to count the number of RS redexes on the starting
term, which should decrease from the point where all steps are erased.
We are now ready to consider the application of these results to higher-order
rewriting.
4.3
Higher-order rewriting à la Nipkow
We can now come back to the case of Nipkow’s rewriting, also called higher-order
rewriting, and understand how Nipkow’s restriction match the previous abstract
result in which the rewrite system S = S1 is made of the two rules β and η, the
second being used as an expansion saturating arrow types.
18
Rules are normalized terms of base type and have a pattern as lefthand side.
Higher-order rewritting (including β) is assumed to be terminating, and higherorder critical pairs (computed by using higher-order matching) are assumed to
be joinable.
Since η is used as an expansion, and has therefore a variable as its lefthand
side, there are no extensions associated with the η-rule and no shallow critical
pairs. For the β-rule, there is no need for extensions: since the rules in R are of
base type, their lefthand side cannot unify with an abstraction. Finally, there
are no shallow pairs either: since rules are in long-βη-normal form, no subterm
of a rule can unify with the lefthand side of β. Therefore,
Corollary 1. Higher-order rewriting is Church-Rosser under Nipkow’s restrictions.
To understand now if these assumptions can be relaxed, let us consider Nipkow’s “counter-examples”.
Example 1. Consider the one rule rewriting system R = {λx.a → λx.b}, where a
and b are two constants of a given base type. a and b are two terms in eta-long
Λ
1
normal form for which a ←→Λ
β (λx.a u) ←→R (λx.b u) ←→β b. But a and b are in
normal form for Nipkow’s rewriting, and therefore a ←→
6 R∪βη b.
This motivates Nipkow’s restriction that the left-hand side is of base type,
and therefore is not an abstraction. Of course, it would be easy to change the
rule into a → b, therefore avoiding the problem, but this cannot be done in
general. Consider for instance the rewrite rule λx. f (x, Z(x)) → λx. g(x, Z(x))
taken from [?]. Removing the abstraction yields f (X, Z(X)) → g(x, Z(X)), a
rule whose left-hand side is not a pattern.
2
This problem can of course be easily solved by adding the rule a → b which
indeed turns out to be the normalized β-extension of the original rule. Note that,
by construction, the left-hand side of the extension is not an abstraction, and
therefore there is no need for β-extensions of β-extensions.
We can of course replace these assumptions by much weaker ones: higherorder rules of arrow type can be considered provided their set is closed under
extensions with the β-rule. Now, it is easy to see that the normalized extension
of a rule λx.l(x) → r is the rule l(x) → @(r, x). Since one abstraction has
been pulled out, a rule can have only finitely many extensions which are easily
computable. The argument that there are no forward pairs to be considered
remains valid. Therefore:
Theorem 3. Given a set R of higher-order rules such that
(i) R ∪ βη−1 is terminating;
(ii) irreducible higher-order critical pairs of R are joinable;
(iii) For each rule λx.l → r ∈ R, R contains its β-extension l → @(r, x) lηβ ;
then, higher-order rewriting with R is Church-Rosser.
It is crucial to notice that we need to prove that the relation −→Rβη ∪ −→β ∪ −→η−1
is terminating, instead of simply −→Rβη and −→β ∪ −→η−1 separately as in
19
Nipkow’s case. While a termination method is known for −→Rβη [?] (and for
−→β ∪ −→η−1 of course) separately, it turns out that no method (up to our
knowledge) is known for proving termination of −→Rβη ∪ −→β ∪ −→η−1 , while
there is one for the relation −→Rβη ∪ −→β ∪ −→η [?]. This is one more motivation for using βη-reductions as presented in Section 2.
Note also the slight improvement obtained by eliminating the reducible higherorder critical pairs from being tested for joinability..
JPJ : Femke, can you care of the end of this paragraph which is in comment
in the LaTeX file? There is also another one before.
4.4
Higher-order rewriting revised
Rewrite rules. Our goal now is to adopt a different definition of higher-order
rewriting by using η-reductions instead of η-expansions. As said in the introduction, this is more natural. We will nevertheless recover the advantages of
η-expansions by having arities for the variables, and therefore variables are now
η-expanded up to the saturation of their arity, as in λxy.X(x, y) for X of arity
two and x, y of arity zero. This η-expanded term is also η-reduced, since X(x)
and X are not terms.
Definition 10. A rewrite rule is an expression of the form Γ ` l → r : σ such
that
(i) l and r are in βη-normal form,
(ii) Γ ` l : σ and Γ ` r : σ,
(iii) l = f (l1 , . . . , ln ) with all li patterns,
(iv) Var(r) ⊆ Var(l)
Rules of functional type are allowed.
We can now derive the Church-Rosser properties of higher-order rewriting
with types and arities, by setting S1 = {β} and S2 = {η}. As we have seen, we
need no shallow pair when lefthand sides are patterns. We need no β-extensions
either, since lefthand sides are now headed by a defined function symbol.
For η, no forward extension is needed, since the righthand side is a variable.
We therefore simply need extensions. It is easy to see that for each rule of
the form @(l, x) → r with x 6∈ Var(l), we need the extension l → λx.r. Note
that λx.r is normalized, and that, again, a rule can have only finitely many
extensions. As a conclusion:
Theorem 4. Given a set R of higher-order rules such that
(i) R ∪ βη is terminating;
(ii) β-irreducible higher-order critical pairs of R are joinable;
(iii) For each rule @(l, x) → r ∈ R such that x 6∈ Var(l), R contains its
η-extension l → λx.r,
then, higher-order rewriting with R is Church-Rosser.
20
Now, we can of course use the results from [?] to prove termination of the
relation −→Rβη ∪ −→βη . And indeed, all sets of rules considered in this paper
can be processed automatically with the tools provided in this paper and the
afore mentionned reference.
Example 2. Types and arities of free variables have to be inferred from their use.
1. The following signature is used to represent the untyped λ-calculus:
app : T × T ⇒ T
abs : (T → T ) ⇒ T
with T a sort representing the set of terms. The β-reduction rule is represented as follows:
app(abs(λx. Z(x)), Z 0 ) → Z(Z 0 )
2. For modelling differentiation, we use the following function symbols:
diff : (R → R) ⇒ (R → R)
sin : R ⇒ R
cos : R ⇒ R
A rewrite rule for differentiation:
diff(λx. sin(x)) → λx. cos(x)
Here we see the advantage of not working with long η-normal forms. In a
HRS format diff(λx. sin(x)) cannot be the left-hand side of a rewrite rule,
because it is not a long η-normal form.
3. For modelling map on finite lists of natural numbers, we use:
nil : natlist
cons : nat × natlist ⇒ natlist
map : natlist × (nat → nat) ⇒ natlist
The rewrite rules for map are as follows:
map(nil, F )
→ nil
map(cons(H, T ), F ) → cons(@(F, H), map(T, F ))
4. Finally an example where we use the following signature:
h:A⇒A
g : A × (A → A → A) ⇒ A
The rewrite rule is as follows:
h(g(X, F )) → g(X, λyz. h(@(F, y, z)))
Using Nipkow’s style, the rule would be instead:
h(g(X, λxy.F (x, y))) → g(X, λyz. h(F (y, z)))
We think that the former is more elegant than the latter. The difference in
the righthand side is superficial: Nipkow allows the writing of F (y, z) for
higher-order variables as a shorthand for @(F, x, y) which we do not, since
this writing is reserved for variables of arity 2.
21
Higher-order unifcation revised Since our framework is made slightly different
from the usual one by using free variables with arities, we need to describe how
higher-order pattern matching and unification is modified. Indeed, the impact of
this change is simply that a free variable X of arity n always comes together with
the (all different) bound variables to which it is applied in order to conform to
the syntax. For example, a solved equation X = t is replaced by an equations of
the form X(x1 , . . . , xn ) = t in which the variables x1 , . . . , xn must be all different
because of the pattern condition, and the variable X does not occur in t. And
therefore, the most general higher-order unifier must substitute the variable X
by the term λx1 · · · xn .t.
Termination of higher-order rewriting Indeed, the solution is given in [?], in the
form of a variation of the higher-order recursive path ordering [?] which allows
to orient terms modulo βη. Besides, the ordering includes β and η-reductions as
well, which fits our purpose.
5
Conclusion
Several problems need investigations.
Polymorphism, as in Jouannaud and Rubio: the same framework applies,
but the unification of patterns reduces to plain unification of types, and then to
usual unification of patterns.
Normalization modulo associativity and commutativity: the abstract ChurchRosser results need to cover this case. On the other hand, pattern unification
modulo AC has been solved already [?]. Higher-order AC-termination has been
investigated also, but only for plain higher-order rewriting. Moving to normal
rewriting should not be a very hard step.
Dependent typing. That’s should be much harder, since almost everything is
to be done. Termination and unification of patterns exist for the whole calculus
of constructions, but need to be adapted or to be improved.
References
1. M. Bezem, J.W. Kop and R. de Vrijer eds. Term Rewriting Systems. Cambridge
Tracts in Theoretical Computer Science 55, Cambridge University Press, 2003.
2. F. Blanqui, J.-P. Jouannaud, and M. Okada. The Calculus of Algebraic Constructions. In Narendran and Rusinowitch, Proc. RTA, LNCS 1631, 1999.
3. F. Blanqui. Inductive Types Revisited. submitted.
4. F. Blanqui. Definitions by rewriting in the Calculus of Constructions, 2003. To
appear in Mathematical Structures in Computer Science.
5. F. Blanqui, J.-P. Jouannaud, and M. Okada. Inductive Data Types. Theoretical
Computer Science 277, 2001.
6. Alexandre Boudet and Evelyne Contejean. AC-Unification of Higher-order Patterns. In Proc. International Conference on Constraionts programming, 1997,
pp:267–281.
22
7. Cristina Borralleras and Albert Rubio. A monotonic, higher-order semantic path
ordering. submitted, 2001.
8. G. Gonthier and J.-J. Lévy and P.-A. Mellies. An abstract Standardisation Theorem. In Proc. 7th IEEE Symp. on Logic in Computer Science, 1992.
9. Gérard Huet. Confluent reductions: abstract properties and applications to term
rewriting systems. Journal of the ACM, 27(4):797–821, October 1980.
10. Jean-Pierre Jouannaud and Hélène Kirchner. Completion of a Set of Rules Modulo
a Set of Equations. In Siam Journal of Computing 15:4(1155–1194), 1984.
11. Jean-Pierre Jouannaud and Mitsuhiro Okada. Abstract data type systems. Theoretical Computer Science, 173(2):349–391, February 1997.
12. Jean-Pierre Jouannaud and Mitsuhiro Okada. Higher-Order Algebraic Specifications. In Annual IEEE Symposium on Logic in Computer Science, Amsterdam,
The Netherlands, 1991. IEEE Comp. Soc. Press.
13. Jean-Pierre Jouannaud and Albert Rubio. The higher-order recursive path ordering. In Giuseppe Longo, editor, Fourteenth Annual IEEE Symposium on Logic in
Computer Science, Trento, Italy, July 1999. IEEE Comp. Soc. Press.
14. Jean-Pierre Jouannaud and Albert Rubio. Higher-order recursive path orderings
“à la carte”. available from the web.
15. Jean-Pierre Jouannaud and Albert Rubio. Higher-order orderings for nmal rewriting. available from the web.
16. Jan Wilhelm Klop. Combinatory Reduction Relations. Mathematical Centre
Tracts 127. Mathematisch Centrum, Amsterdam, 1980.
17. Dallas Lankford.and A. M. Ballantyne. Decision procedures for simple equational
theories with commutative-associative axioms: Complete sets of commutativeassociative reductions. In ”Memo ATP-39”, Dpt of Math. and Comp. Sc., University of Texas, Austin, Texas, 1977.
18. Claude Marché. Normalised Rewriting and Normalised Completion. In Proc. LICS,
pp 394–403, 1994.
19. Richard Mayr and Tobias Nipkow. Higher-order rewrite systems and their confluence. Theoretical Computer Science, 192(1):3–29, February 1998.
20. Tobias Nipkow. Higher-order critical pairs. In 6th IEEE Symp. on Logic in Computer Science, pages 342–349. IEEE Computer Society Press, 1991.
21. Gerald E. Peterson and Mark E. Stickel. Complete sets of reductions for some
equational theories. In JACM 28(2):233–264, 1981.
23

Documents pareils