.. _Conversion-rules: Conversion rules ---------------- Coq has conversion rules that can be used to determine if two terms are equal by definition, or :term:`convertible`. Conversion rules consist of reduction rules and expansion rules. See :ref:`applyingconversionrules`, which describes tactics that apply these conversion rules. α-conversion ~~~~~~~~~~~~ Two terms are :gdef:`α-convertible ` if they are syntactically equal ignoring differences in the names of variables bound within the expression. For example `forall x, x + 0 = x` is α-convertible with `forall y, y + 0 = y`. β-reduction ~~~~~~~~~~~ :gdef:`β-reduction ` reduces a :gdef:`beta-redex`, which is a term in the form `(fun x => t) u`. (Beta-redex is short for "beta-reducible expression", a term from lambda calculus. See `Beta reduction `_ for more background.) Formally, in any :term:`global environment` :math:`E` and :term:`local context` :math:`Γ`, the beta-reduction rule is: .. inference:: Beta -------------- E[Γ] ⊢ ((λx:T.~t)~u)~\triangleright_β~\subst{t}{x}{u} We say that :math:`\subst{t}{x}{u}` is the *β-contraction* of :math:`((λx:T.~t)~u)` and, conversely, that :math:`((λ x:T.~t)~u)` is the *β-expansion* of :math:`\subst{t}{x}{u}`. .. todo: :term:`Calculus of Inductive Constructions` fails to build in CI for some reason :-() Terms of the *Calculus of Inductive Constructions* enjoy some fundamental properties such as confluence, strong normalization, subject reduction. These results are theoretically of great importance but we will not detail them here and refer the interested reader to :cite:`Coq85`. .. _delta-reduction-sect: δ-reduction ~~~~~~~~~~~ :gdef:`δ-reduction ` replaces variables defined in :term:`local contexts ` or :term:`constants ` defined in the :term:`global environment` with their values. :gdef:`Unfolding ` means to replace a constant by its definition. Formally, this is: .. inference:: Delta-Local \WFE{\Gamma} (x:=t:T) ∈ Γ -------------- E[Γ] ⊢ x~\triangleright_Δ~t .. inference:: Delta-Global \WFE{\Gamma} (c:=t:T) ∈ E -------------- E[Γ] ⊢ c~\triangleright_δ~t :term:`Delta-reduction ` only unfolds :term:`constants ` that are marked :gdef:`transparent`. :gdef:`Opaque ` is the opposite of transparent; :term:`delta-reduction` doesn't unfold opaque constants. ι-reduction ~~~~~~~~~~~ A specific conversion rule is associated with the inductive objects in the global environment. We shall give later on (see Section :ref:`Well-formed-inductive-definitions`) the precise rules but it just says that a destructor applied to an object built from a constructor behaves as expected. This reduction is called :gdef:`ι-reduction ` and is more precisely studied in :cite:`Moh93,Wer94`. ζ-reduction ~~~~~~~~~~~ :gdef:`ζ-reduction ` removes :ref:`let-in definitions ` in terms by replacing the defined variable by its value. One way this reduction differs from δ-reduction is that the declaration is removed from the term entirely. Formally, this is: .. inference:: Zeta \WFE{\Gamma} \WTEG{u}{U} \WTE{\Gamma::(x:=u:U)}{t}{T} -------------- E[Γ] ⊢ \letin{x}{u:U}{t}~\triangleright_ζ~\subst{t}{x}{u} .. _eta-expansion-sect: η-expansion ~~~~~~~~~~~ Another important concept is :gdef:`η-expansion `. It is legal to identify any term :math:`t` of functional type :math:`∀ x:T,~U` with its so-called η-expansion .. math:: λx:T.~(t~x) for :math:`x` an arbitrary variable name fresh in :math:`t`. .. note:: We deliberately do not define η-reduction: .. math:: λ x:T.~(t~x)~\not\triangleright_η~t This is because, in general, the type of :math:`t` need not to be convertible to the type of :math:`λ x:T.~(t~x)`. E.g., if we take :math:`f` such that: .. math:: f ~:~ ∀ x:\Type(2),~\Type(1) then .. math:: λ x:\Type(1).~(f~x) ~:~ ∀ x:\Type(1),~\Type(1) We could not allow .. math:: λ x:\Type(1).~(f~x) ~\triangleright_η~ f because the type of the reduced term :math:`∀ x:\Type(2),~\Type(1)` would not be convertible to the type of the original term :math:`∀ x:\Type(1),~\Type(1)`. .. _proof-irrelevance: Proof Irrelevance ~~~~~~~~~~~~~~~~~ It is legal to identify any two terms whose common type is a strict proposition :math:`A : \SProp`. Terms in a strict propositions are therefore called *irrelevant*. .. _convertibility: Convertibility ~~~~~~~~~~~~~~ Let us write :math:`E[Γ] ⊢ t \triangleright u` for the contextual closure of the relation :math:`t` reduces to :math:`u` in the global environment :math:`E` and local context :math:`Γ` with one of the previous reductions β, δ, ι or ζ. We say that two terms :math:`t_1` and :math:`t_2` are *βδιζη-convertible*, or simply :gdef:`convertible`, or :term:`definitionally equal `, in the global environment :math:`E` and local context :math:`Γ` iff there exist terms :math:`u_1` and :math:`u_2` such that :math:`E[Γ] ⊢ t_1 \triangleright … \triangleright u_1` and :math:`E[Γ] ⊢ t_2 \triangleright … \triangleright u_2` and either :math:`u_1` and :math:`u_2` are identical up to irrelevant subterms, or they are convertible up to η-expansion, i.e. :math:`u_1` is :math:`λ x:T.~u_1'` and :math:`u_2 x` is recursively convertible to :math:`u_1'`, or, symmetrically, :math:`u_2` is :math:`λx:T.~u_2'` and :math:`u_1 x` is recursively convertible to :math:`u_2'`. We then write :math:`E[Γ] ⊢ t_1 =_{βδιζη} t_2`. Apart from this we consider two instances of polymorphic and cumulative (see Chapter :ref:`polymorphicuniverses`) inductive types (see below) convertible .. math:: E[Γ] ⊢ t~w_1 … w_m =_{βδιζη} t~w_1' … w_m' if we have subtypings (see below) in both directions, i.e., .. math:: E[Γ] ⊢ t~w_1 … w_m ≤_{βδιζη} t~w_1' … w_m' and .. math:: E[Γ] ⊢ t~w_1' … w_m' ≤_{βδιζη} t~w_1 … w_m. Furthermore, we consider .. math:: E[Γ] ⊢ c~v_1 … v_m =_{βδιζη} c'~v_1' … v_m' convertible if .. math:: E[Γ] ⊢ v_i =_{βδιζη} v_i' and we have that :math:`c` and :math:`c'` are the same constructors of different instances of the same inductive types (differing only in universe levels) such that .. math:: E[Γ] ⊢ c~v_1 … v_m : t~w_1 … w_m and .. math:: E[Γ] ⊢ c'~v_1' … v_m' : t'~ w_1' … w_m ' and we have .. math:: E[Γ] ⊢ t~w_1 … w_m =_{βδιζη} t~w_1' … w_m'. The convertibility relation allows introducing a new typing rule which says that two convertible well-formed types have the same inhabitants.