Documentation

Cslib.Languages.LambdaCalculus.LocallyNameless.Untyped.Basic

λ-calculus #

The untyped λ-calculus, with a locally nameless representation of syntax.

References #

Syntax of locally nameless lambda terms, with free variables over Var.

  • bvar {Var : Type u} : Term Var

    Bound variables that appear under a lambda abstraction, using a de-Bruijn index.

  • fvar {Var : Type u} : VarTerm Var

    Free variables.

  • abs {Var : Type u} : Term VarTerm Var

    Lambda abstraction, introducing a new bound variable.

  • app {Var : Type u} : Term VarTerm VarTerm Var

    Function application.

Instances For

    Variable opening of the ith bound variable.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      theorem Cslib.LambdaCalculus.LocallyNameless.Untyped.Term.openRec_bvar {i : } {Var✝ : Type u_1} {s : Term Var✝} {i' : } :
      openRec i s (bvar i') = if i = i' then s else bvar i'
      theorem Cslib.LambdaCalculus.LocallyNameless.Untyped.Term.openRec_fvar {i : } {Var✝ : Type u_1} {s : Term Var✝} {x : Var✝} :
      openRec i s (fvar x) = fvar x
      theorem Cslib.LambdaCalculus.LocallyNameless.Untyped.Term.openRec_app {i : } {Var✝ : Type u_1} {s l r : Term Var✝} :
      openRec i s (l.app r) = (openRec i s l).app (openRec i s r)
      theorem Cslib.LambdaCalculus.LocallyNameless.Untyped.Term.openRec_abs {i : } {Var✝ : Type u_1} {s M : Term Var✝} :
      openRec i s M.abs = (openRec (i + 1) s M).abs

      Variable opening of the closest binding.

      Equations
      Instances For

        Variable opening of the closest binding.

        Equations
        • One or more equations did not get rendered due to their size.
        Instances For

          Variable closing, replacing a free fvar x with bvar k

          Equations
          • One or more equations did not get rendered due to their size.
          Instances For

            Variable closing of the closest binding.

            Equations
            Instances For

              Variable closing of the closest binding.

              Equations
              • One or more equations did not get rendered due to their size.
              Instances For

                Locally closed terms.

                Instances For

                  Values are irreducible terms.

                  Instances For
                    theorem Cslib.LambdaCalculus.LocallyNameless.Untyped.Term.closeRec_fvar {Var : Type u} [DecidableEq Var] {x : Var} {k : } {x' : Var} :
                    closeRec k x (fvar x') = if x = x' then bvar k else fvar x'
                    theorem Cslib.LambdaCalculus.LocallyNameless.Untyped.Term.closeRec_app {Var : Type u} [DecidableEq Var] {x : Var} {k : } {l r : Term Var} :
                    closeRec k x (l.app r) = (closeRec k x l).app (closeRec k x r)
                    theorem Cslib.LambdaCalculus.LocallyNameless.Untyped.Term.closeRec_abs {Var : Type u} [DecidableEq Var] {x : Var} {k : } {t : Term Var} :
                    closeRec k x t.abs = (closeRec (k + 1) x t).abs
                    theorem Cslib.LambdaCalculus.LocallyNameless.Untyped.Term.subst_fvar {Var : Type u} [DecidableEq Var] {x : Var} {n : Term Var} {x' : Var} :
                    (fvar x')[x:=n] = if x = x' then n else fvar x'
                    theorem Cslib.LambdaCalculus.LocallyNameless.Untyped.Term.subst_app {Var : Type u} [DecidableEq Var] {x : Var} {n l r : Term Var} :
                    (l.app r)[x:=n] = l[x:=n].app (r[x:=n])
                    theorem Cslib.LambdaCalculus.LocallyNameless.Untyped.Term.subst_def {Var : Type u} [DecidableEq Var] (m : Term Var) (x : Var) (n : Term Var) :
                    m.subst x n = m[x:=n]