Additional operations on expr and related types
This file defines basic operations on the types expr, name, declaration, level, environment.
This file is mostly for non-tactics. Tactics should generally be placed in tactic.core.
Tags
expr, name, declaration, level, environment, meta, metaprogramming, tactic
Declarations about binder_info
        Equations
The brackets corresponding to a given binder_info.
Equations
- binder_info.aux_decl.brackets = ("(", ")")
- binder_info.inst_implicit.brackets = ("[", "]")
- binder_info.strict_implicit.brackets = ("{{", "}}")
- binder_info.implicit.brackets = ("{", "}")
- binder_info.default.brackets = ("(", ")")
Declarations about name
        Find the largest prefix n of a name such that f n ≠ none, then replace this prefix
with the value of f n.
Equations
- name.map_prefix f (name.mk_numeral d n') = (f (name.mk_numeral d n')).get_or_else (name.mk_numeral d (name.map_prefix f n'))
- name.map_prefix f (name.mk_string s n') = (f (name.mk_string s n')).get_or_else (name.mk_string s (name.map_prefix f n'))
- name.map_prefix f name.anonymous = name.anonymous
If nm is a simple name (having only one string component) starting with _, then
deinternalize_field nm removes the underscore. Otherwise, it does nothing.
get_nth_prefix nm n removes the last n components from nm
Build a name from components. For example from_components ["foo","bar"] becomes
 `foo.bar
Equations
- name.from_components = from_components_aux name.anonymous
names can contain numeral pieces, which are not legal names
 when typed/passed directly to the parser. We turn an arbitrary
 name into a legal identifier name by turning the numbers to strings.
Append a string to the last component of a name
Equations
- (name.mk_numeral a a_1).append_suffix _x = name.mk_numeral a a_1
- (name.mk_string s n).append_suffix s' = name.mk_string (s ++ s') n
- name.anonymous.append_suffix _x = name.anonymous
Returns the number of characters used to print all the string components of a name, including periods between name segments. Ignores numerical parts of a name.
Checks whether nm has a prefix (including itself) such that P is true
Equations
- name.has_prefix P (name.mk_numeral s nm) = decidable.to_bool (↥(P (name.mk_numeral s nm)) ∨ ↥(name.has_prefix P nm))
- name.has_prefix P (name.mk_string s nm) = decidable.to_bool (↥(P (name.mk_string s nm)) ∨ ↥(name.has_prefix P nm))
- name.has_prefix P name.anonymous = bool.ff
last_string n returns the rightmost component of n, ignoring numeral components.
For example, last_string `a.b.c.33 will return `c.
Equations
- (name.mk_numeral _x n).last_string = n.last_string
- (name.mk_string s _x).last_string = s
- name.anonymous.last_string = "[anonymous]"
Constructs a (non-simple) name from a string.
Example: name.from_string "foo.bar" = `foo.bar
Declarations about level
        Tests whether a universe level is non-zero for all assignments of its variables
l.fold_mvar f folds a function f : name → α → α
over each n : name appearing in a level.mvar n in l.
Declarations about binder
        - name : name
- info : binder_info
- type : expr
The type of binders containing a name, the binding info and the binding type
Turn a binder into a string. Uses expr.to_string for the type.
Converting between expressions and numerals
There are a number of ways to convert between expressions and numerals, depending on the input and output types and whether you want to infer the necessary type classes.
See also the tactics expr.of_nat, expr.of_int, expr.of_rat.
nat.to_pexpr n creates a pexpr that will evaluate to n.
The pexpr does not hold any typing information:
to_expr ``((%%(nat.to_pexpr 5) : ℤ)) will create a native integer numeral (5 : ℤ).
Turns an expression into a natural number, assuming it is only built up from
has_one.one, bit0, bit1, has_zero.zero, nat.zero, and nat.succ.
Turns an expression into a integer, assuming it is only built up from
has_one.one, bit0, bit1, has_zero.zero and a optionally a single has_neg.neg as head.
is_num_eq n1 n2 returns true if n1 and n2 are both numerals with the same numeral structure,
ignoring differences in type and type class arguments.
Declarations about expr
        List of names removed by clean. All these names must resolve to functions defeq id.
Replace any metavariables in the expression with underscores, in preparation for printing
refine ... statements.
If e is a local constant, to_implicit_local_const e changes the binder info of e to
implicit. See also to_implicit_binder, which also changes lambdas and pis.
If e is a local constant, lamda, or pi expression, to_implicit_binder e changes the binder
info of e to implicit. See also to_implicit_local_const, which only changes local constants.
Returns a list of all local constants in an expression (without duplicates).
Returns a name_set of all constants in an expression.
Returns a list of all meta-variables in an expression (without duplicates).
Returns a list of all universe meta-variables in an expression (without duplicates).
Test t contains the specified subexpression e, or a metavariable.
This represents the notion that e "may occur" in t,
possibly after subsequent unification.
Returns a name_set of all constants in an expression starting with a certain prefix.
Returns true if e contains a name n where p n is true.
 Returns true if p name.anonymous is true.
Simplifies the expression t with the specified options.
 The result is (new_e, pr) with the new expression new_e and a proof
 pr : e = new_e.
Definitionally simplifies the expression t with the specified options.
 The result is the simplified expression.
Get the names of the bound variables by a sequence of pis or lambdas.
instantiate_lambdas_or_apps es e instantiates lambdas in e by expressions from es.
If the length of es is larger than the number of lambdas in e,
then the term is applied to the remaining terms.
Also reduces head let-expressions in e, including those after instantiating all lambdas.
This is very similar to expr.substs, but this also reduces head let-expressions.
Get the codomain/target of a pi-type. This definition doesn't instantiate bound variables, and therefore produces a term that is open. See note [open expressions].
Get the body/value of a lambda-expression. This definition doesn't instantiate bound variables, and therefore produces a term that is open. See note [open expressions].
Auxilliary defintion for pi_binders.
 See note [open expressions].
Get the binders and codomain of a pi-type.
 This definition doesn't instantiate bound variables, and therefore produces a term that is open.
 The.tactic get_pi_binders in tactic.core does the same, but also instantiates the
 free variables.
 See note [open expressions].
local_binding_info e returns the binding info of e if e is a local constant.
Otherwise returns binder_info.default.
is_default_local e tests whether e is a local constant with binder info
binder_info.default
has_local_constant e l checks whether local constant l occurs in expression e
Strip-away the context-dependent unique id for the given local const and return: its friendly
name, its binder_info, and its type : expr.
unsafe_cast e freely changes the elab : bool parameter of the passed expr. Mainly used to
access core expr manipulation functions for pexpr-based use, but which are restricted to
expr tt at the site of definition unnecessarily.
DANGER: Unless you know exactly what you are doing, this is probably not the function you are
looking for. For pexpr → expr see tactic.to_expr. For expr → pexpr see to_pexpr.
is_implicitly_included_variable e vs accepts e, an expr.local_const, and a list vs of
   other expr.local_consts. It determines whether e should be considered "available in context"
   as a variable by virtue of the fact that the variables vs have been deemed such.
For example, given variables (n : ℕ) [prime n] [ih : even n], a reference to n implies that
   the typeclass instance prime n should be included, but ih : even n should not.
DANGER: It is possible that for f : expr another expr.local_const, we have
   is_implicitly_included_variable f vs = ff but
   is_implicitly_included_variable f (e :: vs) = tt. This means that one usually wants to
   iteratively add a list of local constants (usually, the variables declared in the local scope)
   which satisfy is_implicitly_included_variable to an initial vs, repeating if any variables
   were added in a particular iteration. The function all_implicitly_included_variables below
   implements this behaviour.
Note that if e ∈ vs then is_implicitly_included_variable e vs = tt.
all_implicitly_included_variables es vs accepts es, a list of expr.local_const, and vs,
   another such list. It returns a list of all variables e in es or vs for which an inclusion
   of the variables in vs into the local context implies that e should also be included. See
   is_implicitly_included_variable e vs for the details.
In particular, those elements of vs are included automatically.
Declarations about environment
        Tests whether a name is declared in the current file. Fixes an error in in_current_file
 which returns tt for the four names quot, quot.mk, quot.lift, quot.ind
Get the full names of all projections of the structure n. Returns none if n is not a
 structure.
Tests whether nm is a generalized inductive type that is not a normal inductive type.
 Note that is_ginductive returns tt even on regular inductive types.
 This returns tt if nm is (part of a) mutually defined inductive type or a nested inductive
 type.
For all declarations d where f d = some x this adds x to the returned list.
Maps f to all declarations in the environment.
Lists all declarations in the environment
Lists all trusted (non-meta) declarations in the environment
Lists the name of all declarations in the environment
Fold a monad over all declarations in the environment.
Filters all declarations in the environment.
Filters all declarations in the environment.
Checks whether s is a prefix of the file where n is declared.
 This is used to check whether n is declared in mathlib, where s is the mathlib directory.
is_eta_expansion
In this section we define the tactic is_eta_expansion which checks whether an expression
  is an eta-expansion of a structure. (not to be confused with eta-expanion for λ).
is_eta_expansion_of args univs l checks whether for all elements (nm, pr) in l we have
 pr = nm.{univs} args.
 Used in is_eta_expansion, where l consists of the projections and the fields of the value we
 want to eta-reduce.
is_eta_expansion_test l checks whether there is a list of expresions args such that for all
 elements (nm, pr) in l we have pr = nm args. If so, returns the last element of args.
 Used in is_eta_expansion, where l consists of the projections and the fields of the value we
 want to eta-reduce.
is_eta_expansion_aux val l checks whether val can be eta-reduced to an expression e.
 Here l is intended to consists of the projections and the fields of val.
 This tactic calls is_eta_expansion_test l, but first removes all proofs from the list l and
 afterward checks whether the resulting expression e unifies with val.
 This last check is necessary, because val and e might have different types.
is_eta_expansion val checks whether there is an expression e such that val is the
 eta-expansion of e.
 With eta-expansion we here mean the eta-expansion of a structure, not of a function.
 For example, the eta-expansion of x : α × β is ⟨x.1, x.2⟩.
 This assumes that val is a fully-applied application of the constructor of a structure.
This is useful to reduce expressions generated by the notation
   { field_1 := _, ..other_structure }
 If other_structure is itself a field of the structure, then the elaborator will insert an
 eta-expanded version of other_structure.
Declarations about declaration
        declaration.update_with_fun f tgt decl
sets the name of the given decl : declaration to tgt, and applies f to the names
of all expr.consts which appear in the value or type of decl.
Checks whether the declaration is declared in the current file.
 This is a simple wrapper around environment.in_current_file'
 Use environment.in_current_file' instead if performance matters.
Checks whether a declaration is automatically generated in the environment. There is no cheap way to check whether a declaration in the namespace of a generalized inductive type is automatically generated, so for now we say that all of them are automatically generated.
Returns true iff d is an automatically-generated or internal declaration.
Returns the list of universe levels of a declaration.
Returns the reducibility_hints field of a defn, and reducibility_hints.opaque otherwise
pretty-prints a declaration object.