feat(frontends/lean): add 'reducible' modifier for controlling which
definitions are unfolded during elaboration Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
This commit is contained in:
parent
baf4c01de8
commit
08ccd58eb6
33 changed files with 360 additions and 128 deletions
|
@ -6,10 +6,10 @@ namespace function
|
|||
section
|
||||
parameters {A : Type} {B : Type} {C : Type} {D : Type} {E : Type}
|
||||
|
||||
definition compose (f : B → C) (g : A → B) : A → C :=
|
||||
definition compose [reducible] (f : B → C) (g : A → B) : A → C :=
|
||||
λx, f (g x)
|
||||
|
||||
definition id (a : A) : A :=
|
||||
definition id [reducible] (a : A) : A :=
|
||||
a
|
||||
|
||||
definition on_fun (f : B → B → C) (g : A → B) : A → A → C :=
|
||||
|
|
|
@ -143,7 +143,6 @@ or.elim le_total
|
|||
... = pr2 a + 0 : add_zero_right⁻¹
|
||||
... = pr2 a + pr1 (proj a) : {(proj_le_pr1 H)⁻¹})
|
||||
|
||||
opaque add sub le
|
||||
theorem proj_congr {a b : ℕ × ℕ} (H : rel a b) : proj a = proj b :=
|
||||
have special : ∀a b, pr2 a ≤ pr1 a → rel a b → proj a = proj b, from
|
||||
take a b,
|
||||
|
@ -207,7 +206,7 @@ exists_intro (pr1 (rep a))
|
|||
theorem has_decidable_eq [instance] [protected] : decidable_eq ℤ :=
|
||||
take a b : ℤ, _
|
||||
|
||||
opaque int
|
||||
reducible [off] int
|
||||
definition of_nat [coercion] (n : ℕ) : ℤ := psub (pair n 0)
|
||||
|
||||
theorem eq_zero_intro (n : ℕ) : psub (pair n n) = 0 :=
|
||||
|
@ -301,7 +300,7 @@ have H5 : n + m = 0, from
|
|||
|
||||
---reverse equalities
|
||||
|
||||
transparent int
|
||||
reducible int
|
||||
|
||||
theorem cases (a : ℤ) : (∃n : ℕ, a = of_nat n) ∨ (∃n : ℕ, a = - n) :=
|
||||
have Hrep : proj (rep a) = rep a, from @idempotent_image_fix _ proj proj_idempotent a,
|
||||
|
@ -324,7 +323,7 @@ or.imp_or (or.swap (proj_zero_or (rep a)))
|
|||
... = -(psub (pair (pr2 (rep a)) 0)) : by simp
|
||||
... = -(of_nat (pr2 (rep a))) : rfl))
|
||||
|
||||
opaque int
|
||||
reducible [off] int
|
||||
|
||||
---rename to by_cases in Lean 0.2 (for now using this to avoid name clash)
|
||||
theorem int_by_cases {P : ℤ → Prop} (a : ℤ) (H1 : ∀n : ℕ, P (of_nat n)) (H2 : ∀n : ℕ, P (-n)) :
|
||||
|
@ -373,7 +372,7 @@ have H4 : n + m = 0, from
|
|||
add_eq_zero H4
|
||||
|
||||
--some of these had to be transparent for theorem cases
|
||||
opaque psub proj
|
||||
reducible [off] psub proj
|
||||
|
||||
-- ## add
|
||||
|
||||
|
@ -466,8 +465,6 @@ rfl
|
|||
theorem add_neg_left (a b : ℤ) : -a + b = b - a :=
|
||||
add_comm (-a) b
|
||||
|
||||
-- opaque_hint (hiding int.sub)
|
||||
|
||||
theorem sub_neg_right (a b : ℤ) : a - (-b) = a + b :=
|
||||
neg_neg b ▸ eq.refl (a - (-b))
|
||||
|
||||
|
|
|
@ -133,7 +133,6 @@ discriminate
|
|||
|
||||
-- ### interaction with neg and sub
|
||||
|
||||
opaque le neg
|
||||
theorem le_neg {a b : ℤ} (H : a ≤ b) : -b ≤ -a :=
|
||||
obtain (n : ℕ) (Hn : a + n = b), from le_elim H,
|
||||
have H2 : b - n = a, from add_imp_sub_right Hn,
|
||||
|
|
|
@ -72,7 +72,7 @@ theorem pred_zero : pred 0 = 0
|
|||
|
||||
theorem pred_succ {n : ℕ} : pred (succ n) = n
|
||||
|
||||
opaque pred
|
||||
reducible [off] pred
|
||||
|
||||
theorem zero_or_succ_pred (n : ℕ) : n = 0 ∨ n = succ (pred n) :=
|
||||
induction_on n
|
||||
|
@ -154,7 +154,7 @@ theorem add_zero_right {n : ℕ} : n + 0 = n
|
|||
|
||||
theorem add_succ_right {n m : ℕ} : n + succ m = succ (n + m)
|
||||
|
||||
opaque add
|
||||
reducible [off] add
|
||||
|
||||
theorem add_zero_left {n : ℕ} : 0 + n = n :=
|
||||
induction_on n
|
||||
|
@ -265,7 +265,7 @@ theorem mul_zero_right {n : ℕ} : n * 0 = 0
|
|||
|
||||
theorem mul_succ_right {n m : ℕ} : n * succ m = n * m + n
|
||||
|
||||
opaque mul
|
||||
reducible [off] mul
|
||||
|
||||
-- ### commutativity, distributivity, associativity, identity
|
||||
|
||||
|
|
|
@ -305,7 +305,6 @@ case_zero_pos n (by simp)
|
|||
|
||||
-- add_rewrite mod_mul_self_right
|
||||
|
||||
opaque add mul
|
||||
theorem mod_mul_self_left {m n : ℕ} : (m * n) mod m = 0 :=
|
||||
mul_comm ▸ mod_mul_self_right
|
||||
|
||||
|
|
|
@ -30,7 +30,7 @@ exists_intro k H
|
|||
theorem le_elim {n m : ℕ} (H : n ≤ m) : ∃k, n + k = m :=
|
||||
H
|
||||
|
||||
opaque le
|
||||
reducible [off] le
|
||||
|
||||
-- ### partial order (totality is part of less than)
|
||||
|
||||
|
|
|
@ -26,7 +26,7 @@ theorem sub_zero_right {n : ℕ} : n - 0 = n
|
|||
|
||||
theorem sub_succ_right {n m : ℕ} : n - succ m = pred (n - m)
|
||||
|
||||
opaque sub
|
||||
reducible [off] sub
|
||||
|
||||
theorem sub_zero_left {n : ℕ} : 0 - n = 0 :=
|
||||
induction_on n sub_zero_right
|
||||
|
@ -67,7 +67,6 @@ induction_on k
|
|||
... = (n + l) - (m + l) : sub_succ_succ
|
||||
... = n - m : IH)
|
||||
|
||||
opaque add mul le
|
||||
theorem sub_add_add_left {k n m : ℕ} : (k + n) - (k + m) = n - m :=
|
||||
add_comm ▸ add_comm ▸ sub_add_add_right
|
||||
|
||||
|
|
|
@ -193,8 +193,6 @@ theorem comp_quotient_map_binary_refl {A B : Type} {R : A → A → Prop} (Hrefl
|
|||
(a b : A) : quotient_map_binary Q f (abs a) (abs b) = abs (f a b) :=
|
||||
comp_quotient_map_binary Q H (Hrefl a) (Hrefl b)
|
||||
|
||||
opaque rec rec_constant rec_binary quotient_map quotient_map_binary
|
||||
|
||||
-- image
|
||||
-- -----
|
||||
definition image {A B : Type} (f : A → B) := subtype (fun b, ∃a, f a = b)
|
||||
|
|
|
@ -14,6 +14,7 @@ axiom funext : ∀ {A : Type} {B : A → Type} {f g : Π x, B x} (H : ∀ x, f x
|
|||
namespace function
|
||||
section
|
||||
parameters {A B C D: Type}
|
||||
|
||||
theorem compose_assoc (f : C → D) (g : B → C) (h : A → B) : (f ∘ g) ∘ h = f ∘ (g ∘ h) :=
|
||||
funext (take x, rfl)
|
||||
|
||||
|
|
|
@ -48,8 +48,6 @@ eq_to_heq (eq.refl a)
|
|||
theorem heqt_elim {a : Prop} (H : a == true) : a :=
|
||||
eq_true_elim (heq_to_eq H)
|
||||
|
||||
opaque cast
|
||||
|
||||
theorem hsubst {A B : Type} {a : A} {b : B} {P : ∀ (T : Type), T → Prop} (H1 : a == b)
|
||||
(H2 : P A a) : P B b :=
|
||||
have Haux1 : ∀ H : A = A, P A (cast H a), from
|
||||
|
|
|
@ -4,7 +4,7 @@
|
|||
|
||||
import general_notation
|
||||
|
||||
definition Prop := Type.{0}
|
||||
definition Prop [reducible] := Type.{0}
|
||||
|
||||
|
||||
-- implication
|
||||
|
|
|
@ -8,7 +8,7 @@
|
|||
(require 'rx)
|
||||
|
||||
(defconst lean-keywords
|
||||
'("import" "opaque" "transparent" "tactic_hint" "definition" "renaming"
|
||||
'("import" "reducible" "tactic_hint" "definition" "renaming"
|
||||
"inline" "hiding" "exposing" "parameter" "parameters" "begin" "proof" "qed" "conjecture"
|
||||
"hypothesis" "lemma" "corollary" "variable" "variables" "print" "theorem"
|
||||
"context" "open" "as" "export" "axiom" "inductive" "with" "structure" "universe" "alias" "help" "environment"
|
||||
|
@ -70,8 +70,8 @@
|
|||
;; place holder
|
||||
(,(rx symbol-start "_" symbol-end) . 'font-lock-preprocessor-face)
|
||||
;; modifiers
|
||||
(,(rx (or "\[persistent\]" "\[notation\]" "\[visible\]" "\[protected\]" "\[private\]"
|
||||
"\[instance\]" "\[class\]" "\[coercion\]" "\[inline\]")) . 'font-lock-doc-face)
|
||||
(,(rx (or "\[persistent\]" "\[notation\]" "\[opaque\]" "\[visible\]" "\[protected\]" "\[private\]"
|
||||
"\[instance\]" "\[class\]" "\[coercion\]" "\[off\]" "\[none\]" "\[on\]")) . 'font-lock-doc-face)
|
||||
;; tactics
|
||||
(,(rx symbol-start
|
||||
(or "\\b.*_tac" "Cond" "or_else" "then" "try" "when" "assumption" "apply" "back" "beta" "done" "exact" "repeat")
|
||||
|
|
|
@ -16,7 +16,7 @@ Author: Leonardo de Moura
|
|||
#include "library/protected.h"
|
||||
#include "library/locals.h"
|
||||
#include "library/coercion.h"
|
||||
#include "library/opaque_hints.h"
|
||||
#include "library/reducible.h"
|
||||
#include "frontends/lean/util.h"
|
||||
#include "frontends/lean/parser.h"
|
||||
#include "frontends/lean/calc.h"
|
||||
|
@ -45,6 +45,9 @@ static name g_exposing("exposing");
|
|||
static name g_renaming("renaming");
|
||||
static name g_as("as");
|
||||
static name g_colon(":");
|
||||
static name g_on("[on]");
|
||||
static name g_off("[off]");
|
||||
static name g_none("[none]");
|
||||
|
||||
environment print_cmd(parser & p) {
|
||||
if (p.curr() == scanner::token_kind::String) {
|
||||
|
@ -105,7 +108,7 @@ environment check_cmd(parser & p) {
|
|||
level_param_names ls = to_level_param_names(collect_univ_params(e));
|
||||
level_param_names new_ls;
|
||||
std::tie(e, new_ls) = p.elaborate_relaxed(e, ctx);
|
||||
auto tc = mk_type_checker_with_hints(p.env(), p.mk_ngen(), true);
|
||||
auto tc = mk_type_checker(p.env(), p.mk_ngen(), true);
|
||||
expr type = tc->check(e, append(ls, new_ls)).first;
|
||||
auto reg = p.regular_stream();
|
||||
formatter const & fmt = reg.get_formatter();
|
||||
|
@ -307,30 +310,40 @@ environment coercion_cmd(parser & p) {
|
|||
}
|
||||
}
|
||||
|
||||
environment opaque_cmd_core(parser & p, bool hiding) {
|
||||
static void parse_reducible_modifiers(parser & p, reducible_status & status, bool & persistent) {
|
||||
while (true) {
|
||||
if (parse_persistent(p, persistent)) {
|
||||
} else if (p.curr_is_token_or_id(g_on)) {
|
||||
p.next();
|
||||
status = reducible_status::On;
|
||||
} else if (p.curr_is_token_or_id(g_off)) {
|
||||
p.next();
|
||||
status = reducible_status::Off;
|
||||
} else if (p.curr_is_token_or_id(g_none)) {
|
||||
p.next();
|
||||
status = reducible_status::None;
|
||||
} else {
|
||||
break;
|
||||
}
|
||||
}
|
||||
}
|
||||
|
||||
environment reducible_cmd(parser & p) {
|
||||
environment env = p.env();
|
||||
reducible_status status = reducible_status::On;
|
||||
bool persistent = false;
|
||||
parse_reducible_modifiers(p, status, persistent);
|
||||
bool found = false;
|
||||
while (p.curr_is_identifier()) {
|
||||
name c = p.check_constant_next("invalid 'opaque/transparent' command, constant expected");
|
||||
name c = p.check_constant_next("invalid 'reducible' command, constant expected");
|
||||
found = true;
|
||||
if (hiding)
|
||||
env = hide_definition(env, c);
|
||||
else
|
||||
env = expose_definition(env, c);
|
||||
env = set_reducible(env, c, status, persistent);
|
||||
}
|
||||
if (!found)
|
||||
throw exception("invalid empty 'opaque/transparent' command");
|
||||
throw exception("invalid empty 'reducible' command");
|
||||
return env;
|
||||
}
|
||||
|
||||
environment opaque_cmd(parser & p) {
|
||||
return opaque_cmd_core(p, true);
|
||||
}
|
||||
|
||||
environment transparent_cmd(parser & p) {
|
||||
return opaque_cmd_core(p, false);
|
||||
}
|
||||
|
||||
environment erase_cache_cmd(parser & p) {
|
||||
name n = p.check_id_next("invalid #erase_cache command, identifier expected");
|
||||
p.erase_cached_definition(n);
|
||||
|
@ -350,8 +363,7 @@ cmd_table init_cmd_table() {
|
|||
add_cmd(r, cmd_info("end", "close the current namespace/section", end_scoped_cmd));
|
||||
add_cmd(r, cmd_info("check", "type check given expression, and display its type", check_cmd));
|
||||
add_cmd(r, cmd_info("coercion", "add a new coercion", coercion_cmd));
|
||||
add_cmd(r, cmd_info("opaque", "mark constants as opaque for the elaborator/unifier", opaque_cmd));
|
||||
add_cmd(r, cmd_info("transparent", "mark constants as transparent for the elaborator/unifier", transparent_cmd));
|
||||
add_cmd(r, cmd_info("reducible", "mark definitions as reducible/irreducible for automation", reducible_cmd));
|
||||
add_cmd(r, cmd_info("#erase_cache", "erase cached definition (for debugging purposes)", erase_cache_cmd));
|
||||
|
||||
register_decl_cmds(r);
|
||||
|
|
|
@ -9,7 +9,7 @@ Author: Leonardo de Moura
|
|||
#include "kernel/instantiate.h"
|
||||
#include "library/scoped_ext.h"
|
||||
#include "library/kernel_serializer.h"
|
||||
#include "library/opaque_hints.h"
|
||||
#include "library/reducible.h"
|
||||
#include "frontends/lean/parser.h"
|
||||
#include "frontends/lean/util.h"
|
||||
#include "frontends/lean/tactic_hint.h"
|
||||
|
@ -102,7 +102,7 @@ environment add_instance(environment const & env, name const & n, bool persisten
|
|||
declaration d = env.get(n);
|
||||
expr type = d.get_type();
|
||||
name_generator ngen(g_tmp_prefix);
|
||||
auto tc = mk_type_checker_with_hints(env, ngen, false);
|
||||
auto tc = mk_type_checker(env, ngen, false);
|
||||
while (true) {
|
||||
type = tc->whnf(type).first;
|
||||
if (!is_pi(type))
|
||||
|
|
|
@ -15,6 +15,7 @@ Author: Leonardo de Moura
|
|||
#include "library/placeholder.h"
|
||||
#include "library/locals.h"
|
||||
#include "library/explicit.h"
|
||||
#include "library/reducible.h"
|
||||
#include "library/coercion.h"
|
||||
#include "frontends/lean/parser.h"
|
||||
#include "frontends/lean/util.h"
|
||||
|
@ -30,6 +31,7 @@ static name g_protected("[protected]");
|
|||
static name g_instance("[instance]");
|
||||
static name g_coercion("[coercion]");
|
||||
static name g_opaque("[opaque]");
|
||||
static name g_reducible("[reducible]");
|
||||
|
||||
environment universe_cmd(parser & p) {
|
||||
name n = p.check_id_next("invalid universe declaration, identifier expected");
|
||||
|
@ -165,12 +167,14 @@ struct decl_modifiers {
|
|||
bool m_is_opaque;
|
||||
bool m_is_instance;
|
||||
bool m_is_coercion;
|
||||
bool m_is_reducible;
|
||||
decl_modifiers() {
|
||||
m_is_private = false;
|
||||
m_is_protected = false;
|
||||
m_is_opaque = true;
|
||||
m_is_instance = false;
|
||||
m_is_coercion = false;
|
||||
m_is_reducible = false;
|
||||
}
|
||||
|
||||
void parse(parser & p) {
|
||||
|
@ -190,6 +194,9 @@ struct decl_modifiers {
|
|||
} else if (p.curr_is_token(g_coercion)) {
|
||||
m_is_coercion = true;
|
||||
p.next();
|
||||
} else if (p.curr_is_token(g_reducible)) {
|
||||
m_is_reducible = true;
|
||||
p.next();
|
||||
} else {
|
||||
break;
|
||||
}
|
||||
|
@ -359,6 +366,8 @@ environment definition_cmd_core(parser & p, bool is_theorem) {
|
|||
env = add_coercion(env, real_n, p.ios());
|
||||
if (modifiers.m_is_protected)
|
||||
env = add_protected(env, real_n);
|
||||
if (modifiers.m_is_reducible)
|
||||
env = set_reducible(env, real_n, reducible_status::On);
|
||||
return env;
|
||||
}
|
||||
environment definition_cmd(parser & p) {
|
||||
|
|
|
@ -24,7 +24,7 @@ Author: Leonardo de Moura
|
|||
#include "library/choice.h"
|
||||
#include "library/explicit.h"
|
||||
#include "library/unifier.h"
|
||||
#include "library/opaque_hints.h"
|
||||
#include "library/reducible.h"
|
||||
#include "library/locals.h"
|
||||
#include "library/let.h"
|
||||
#include "library/deep_copy.h"
|
||||
|
@ -162,8 +162,8 @@ public:
|
|||
m_unifier_config(env.m_ios.get_options(), true /* use exceptions */, true /* discard */) {
|
||||
m_relax_main_opaque = false;
|
||||
m_no_info = false;
|
||||
m_tc[0] = mk_type_checker_with_hints(env.m_env, m_ngen.mk_child(), false);
|
||||
m_tc[1] = mk_type_checker_with_hints(env.m_env, m_ngen.mk_child(), true);
|
||||
m_tc[0] = mk_type_checker(env.m_env, m_ngen.mk_child(), false);
|
||||
m_tc[1] = mk_type_checker(env.m_env, m_ngen.mk_child(), true);
|
||||
}
|
||||
|
||||
environment const & env() const { return m_env.m_env; }
|
||||
|
@ -983,7 +983,7 @@ public:
|
|||
bool relax_main_opaque) {
|
||||
m_context.set_ctx(ctx);
|
||||
m_full_context.set_ctx(ctx);
|
||||
flet<bool> set_relax(m_relax_main_opaque, relax_main_opaque && !get_hide_main_opaque(env()));
|
||||
flet<bool> set_relax(m_relax_main_opaque, relax_main_opaque);
|
||||
constraint_seq cs;
|
||||
expr r = visit(e, cs);
|
||||
if (_ensure_type)
|
||||
|
@ -1002,7 +1002,7 @@ public:
|
|||
constraint_seq t_cs;
|
||||
expr r_t = ensure_type(visit(t, t_cs), t_cs);
|
||||
// Opaque definitions in the main module may treat other opaque definitions (in the main module) as transparent.
|
||||
flet<bool> set_relax(m_relax_main_opaque, is_opaque && !get_hide_main_opaque(env()));
|
||||
flet<bool> set_relax(m_relax_main_opaque, is_opaque);
|
||||
constraint_seq v_cs;
|
||||
expr r_v = visit(v, v_cs);
|
||||
expr r_v_type = infer_type(r_v, v_cs);
|
||||
|
|
|
@ -19,7 +19,7 @@ Author: Leonardo de Moura
|
|||
#include "library/aliases.h"
|
||||
#include "library/protected.h"
|
||||
#include "library/explicit.h"
|
||||
#include "library/opaque_hints.h"
|
||||
#include "library/reducible.h"
|
||||
#include "frontends/lean/decl_cmds.h"
|
||||
#include "frontends/lean/util.h"
|
||||
#include "frontends/lean/class.h"
|
||||
|
@ -105,7 +105,7 @@ struct inductive_cmd_fn {
|
|||
m_u = mk_global_univ(u_name);
|
||||
m_infer_result_universe = false;
|
||||
m_namespace = get_namespace(m_env);
|
||||
m_tc = mk_type_checker_with_hints(m_env, m_p.mk_ngen(), false);
|
||||
m_tc = mk_type_checker(m_env, m_p.mk_ngen(), false);
|
||||
}
|
||||
|
||||
[[ noreturn ]] void throw_error(char const * error_msg) { throw parser_error(error_msg, m_pos); }
|
||||
|
@ -451,7 +451,7 @@ struct inductive_cmd_fn {
|
|||
std::tie(d_name, d_type, std::ignore) = d;
|
||||
m_env = m_env.add(check(m_env, mk_var_decl(d_name, ls, d_type)));
|
||||
}
|
||||
m_tc = mk_type_checker_with_hints(m_env, m_p.mk_ngen(), false);
|
||||
m_tc = mk_type_checker(m_env, m_p.mk_ngen(), false);
|
||||
}
|
||||
|
||||
/** \brief Traverse the introduction rule type and collect the universes where non-parameters reside in \c r_lvls.
|
||||
|
|
|
@ -8,7 +8,7 @@ Author: Leonardo de Moura
|
|||
#include "kernel/instantiate.h"
|
||||
#include "kernel/abstract.h"
|
||||
#include "library/unifier.h"
|
||||
#include "library/opaque_hints.h"
|
||||
#include "library/reducible.h"
|
||||
#include "library/metavar_closure.h"
|
||||
#include "library/error_handling/error_handling.h"
|
||||
#include "frontends/lean/util.h"
|
||||
|
@ -30,7 +30,7 @@ struct placeholder_context {
|
|||
name const & prefix, bool relax, bool use_local_instances):
|
||||
m_ios(ios),
|
||||
m_ngen(prefix),
|
||||
m_tc(mk_type_checker_with_hints(env, m_ngen.mk_child(), relax)),
|
||||
m_tc(mk_type_checker(env, m_ngen.mk_child(), relax)),
|
||||
m_ctx(m_ngen.next(), ctx),
|
||||
m_relax(relax),
|
||||
m_use_local_instances(use_local_instances) {
|
||||
|
|
|
@ -5,7 +5,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
|
|||
Author: Leonardo de Moura
|
||||
*/
|
||||
#include "library/unifier.h"
|
||||
#include "library/opaque_hints.h"
|
||||
#include "library/reducible.h"
|
||||
#include "library/metavar_closure.h"
|
||||
#include "frontends/lean/util.h"
|
||||
#include "frontends/lean/local_context.h"
|
||||
|
@ -17,7 +17,7 @@ constraint mk_proof_qed_cnstr(environment const & env, expr const & m, expr cons
|
|||
auto choice_fn = [=](expr const & meta, expr const & meta_type, substitution const & s,
|
||||
name_generator const & _ngen) {
|
||||
name_generator ngen(_ngen);
|
||||
type_checker_ptr tc(mk_type_checker_with_hints(env, ngen.mk_child(), relax));
|
||||
type_checker_ptr tc(mk_type_checker(env, ngen.mk_child(), relax));
|
||||
pair<expr, constraint_seq> tcs = tc->infer(e);
|
||||
expr e_type = tcs.first;
|
||||
justification new_j = mk_type_mismatch_jst(e, e_type, meta_type);
|
||||
|
|
|
@ -17,7 +17,7 @@ Author: Leonardo de Moura
|
|||
#include "library/aliases.h"
|
||||
#include "library/type_util.h"
|
||||
#include "library/unifier.h"
|
||||
#include "library/opaque_hints.h"
|
||||
#include "library/reducible.h"
|
||||
#include "library/scoped_ext.h"
|
||||
#include "library/tactic/goal.h"
|
||||
#include "frontends/lean/server.h"
|
||||
|
@ -769,7 +769,7 @@ void server::find_goal_matches(unsigned line_num, unsigned col_num, std::string
|
|||
environment const & env = env_opts->first;
|
||||
options const & opts = env_opts->second;
|
||||
name_generator ngen(g_tmp_prefix);
|
||||
std::unique_ptr<type_checker> tc = mk_type_checker_with_hints(env, ngen, true);
|
||||
std::unique_ptr<type_checker> tc = mk_type_checker(env, ngen, true);
|
||||
if (auto meta = m_file->infom().get_meta_at(line_num, col_num)) {
|
||||
if (is_meta(*meta)) {
|
||||
if (auto type = m_file->infom().get_type_at(line_num, col_num)) {
|
||||
|
|
|
@ -15,7 +15,7 @@ Author: Leonardo de Moura
|
|||
#include "library/scoped_ext.h"
|
||||
#include "library/placeholder.h"
|
||||
#include "library/locals.h"
|
||||
#include "library/opaque_hints.h"
|
||||
#include "library/reducible.h"
|
||||
#include "frontends/lean/parser.h"
|
||||
#include "frontends/lean/util.h"
|
||||
#include "frontends/lean/decl_cmds.h"
|
||||
|
@ -235,7 +235,7 @@ struct structure_cmd_fn {
|
|||
This information is used to compute the resultant universe level for the inductive datatype declaration.
|
||||
*/
|
||||
void accumulate_levels(expr intro_type, unsigned num_params, buffer<level> & r_lvls) {
|
||||
auto tc = mk_type_checker_with_hints(m_env, m_p.mk_ngen(), false);
|
||||
auto tc = mk_type_checker(m_env, m_p.mk_ngen(), false);
|
||||
unsigned i = 0;
|
||||
while (is_pi(intro_type)) {
|
||||
if (i >= num_params) {
|
||||
|
|
|
@ -78,8 +78,8 @@ token_table init_token_table() {
|
|||
{"+", g_plus_prec}, {g_cup, g_cup_prec}, {"->", g_arrow_prec}, {nullptr, 0}};
|
||||
|
||||
char const * commands[] = {"theorem", "axiom", "variable", "definition", "coercion",
|
||||
"variables", "[persistent]", "[private]", "[protected]", "[inline]", "[visible]", "[instance]",
|
||||
"[class]", "[coercion]", "[opaque]", "abbreviation", "opaque", "transparent",
|
||||
"variables", "[persistent]", "[private]", "[protected]", "[visible]", "[instance]",
|
||||
"[off]", "[on]", "[none]", "[class]", "[coercion]", "[opaque]", "[reducible]", "reducible",
|
||||
"evaluate", "check", "print", "end", "namespace", "section", "import",
|
||||
"inductive", "record", "renaming", "extends", "structure", "module", "universe",
|
||||
"precedence", "infixl", "infixr", "infix", "postfix", "prefix", "notation", "context",
|
||||
|
|
|
@ -6,9 +6,9 @@ add_library(library deep_copy.cpp expr_lt.cpp io_state.cpp occurs.cpp
|
|||
update_declaration.cpp choice.cpp scoped_ext.cpp locals.cpp
|
||||
standard_kernel.cpp sorry.cpp
|
||||
unifier.cpp unifier_plugin.cpp inductive_unifier_plugin.cpp
|
||||
explicit.cpp num.cpp string.cpp opaque_hints.cpp head_map.cpp
|
||||
match.cpp definition_cache.cpp declaration_index.cpp
|
||||
explicit.cpp num.cpp string.cpp head_map.cpp match.cpp
|
||||
definition_cache.cpp declaration_index.cpp
|
||||
print.cpp annotation.cpp typed_expr.cpp let.cpp type_util.cpp
|
||||
protected.cpp metavar_closure.cpp)
|
||||
protected.cpp metavar_closure.cpp reducible.cpp)
|
||||
|
||||
target_link_libraries(library ${LEAN_LIBS})
|
||||
|
|
|
@ -29,7 +29,7 @@ Author: Leonardo de Moura
|
|||
#include "library/kernel_bindings.h"
|
||||
#include "library/normalize.h"
|
||||
#include "library/module.h"
|
||||
#include "library/opaque_hints.h"
|
||||
#include "library/reducible.h"
|
||||
#include "library/print.h"
|
||||
|
||||
// Lua Bindings for the Kernel classes. We do not include the Lua
|
||||
|
@ -1870,13 +1870,13 @@ static int mk_type_checker_with_hints(lua_State * L) {
|
|||
environment const & env = to_environment(L, 1);
|
||||
int nargs = lua_gettop(L);
|
||||
if (nargs == 1) {
|
||||
return push_type_checker_ref(L, mk_type_checker_with_hints(env, name_generator(g_tmp_prefix), false));
|
||||
return push_type_checker_ref(L, mk_type_checker(env, name_generator(g_tmp_prefix), false));
|
||||
} else if (nargs == 2 && lua_isboolean(L, 2)) {
|
||||
return push_type_checker_ref(L, mk_type_checker_with_hints(env, name_generator(g_tmp_prefix), lua_toboolean(L, 2)));
|
||||
return push_type_checker_ref(L, mk_type_checker(env, name_generator(g_tmp_prefix), lua_toboolean(L, 2)));
|
||||
} else if (nargs == 2) {
|
||||
return push_type_checker_ref(L, mk_type_checker_with_hints(env, to_name_generator(L, 2), false));
|
||||
return push_type_checker_ref(L, mk_type_checker(env, to_name_generator(L, 2), false));
|
||||
} else {
|
||||
return push_type_checker_ref(L, mk_type_checker_with_hints(env, to_name_generator(L, 2), lua_toboolean(L, 3)));
|
||||
return push_type_checker_ref(L, mk_type_checker(env, to_name_generator(L, 2), lua_toboolean(L, 3)));
|
||||
}
|
||||
}
|
||||
|
||||
|
|
|
@ -7,12 +7,13 @@ Author: Leonardo de Moura
|
|||
#include "util/sstream.h"
|
||||
#include "kernel/environment.h"
|
||||
#include "kernel/type_checker.h"
|
||||
#include "library/module.h"
|
||||
|
||||
namespace lean {
|
||||
struct opaque_hints_ext : public environment_extension {
|
||||
name_set m_extra_opaque;
|
||||
bool m_hide_module;
|
||||
opaque_hints_ext():m_hide_module(false) {}
|
||||
name_set m_opaque;
|
||||
name_set m_transparent;
|
||||
opaque_hints_ext() {}
|
||||
};
|
||||
|
||||
struct opaque_hints_ext_reg {
|
||||
|
@ -34,32 +35,39 @@ static void check_definition(environment const & env, name const & n) {
|
|||
environment hide_definition(environment const & env, name const & n) {
|
||||
check_definition(env, n);
|
||||
auto ext = get_extension(env);
|
||||
ext.m_extra_opaque.insert(n);
|
||||
ext.m_opaque.insert(n);
|
||||
ext.m_transparent.erase(n);
|
||||
return update(env, ext);
|
||||
}
|
||||
environment expose_definition(environment const & env, name const & n) {
|
||||
check_definition(env, n);
|
||||
auto ext = get_extension(env);
|
||||
if (!ext.m_extra_opaque.contains(n))
|
||||
throw exception(sstream() << "invalid 'exposing' opaque/transparent, '" << n << "' is not in the 'extra opaque' set");
|
||||
ext.m_extra_opaque.erase(n);
|
||||
ext.m_opaque.erase(n);
|
||||
ext.m_transparent.insert(n);
|
||||
return update(env, ext);
|
||||
}
|
||||
environment set_hide_main_opaque(environment const & env, bool f) {
|
||||
auto ext = get_extension(env);
|
||||
ext.m_hide_module = f;
|
||||
return update(env, ext);
|
||||
}
|
||||
bool get_hide_main_opaque(environment const & env) {
|
||||
return get_extension(env).m_hide_module;
|
||||
bool is_exposed(environment const & env, name const & n) {
|
||||
auto const & ext = get_extension(env);
|
||||
return ext.m_transparent.contains(n);
|
||||
}
|
||||
std::unique_ptr<type_checker> mk_type_checker_with_hints(environment const & env, name_generator const & ngen,
|
||||
bool relax_main_opaque) {
|
||||
bool relax_main_opaque, bool only_main_transparent) {
|
||||
auto const & ext = get_extension(env);
|
||||
name_set extra_opaque = ext.m_extra_opaque;
|
||||
if (only_main_transparent) {
|
||||
name_set extra_opaque = ext.m_opaque;
|
||||
name_set extra_transparent = ext.m_transparent;
|
||||
extra_opaque_pred pred([=](name const & n) {
|
||||
return
|
||||
(!module::is_definition(env, n) || extra_opaque.contains(n)) &&
|
||||
(!extra_transparent.contains(n));
|
||||
});
|
||||
return std::unique_ptr<type_checker>(new type_checker(env, ngen, mk_default_converter(env, relax_main_opaque,
|
||||
true, pred)));
|
||||
} else {
|
||||
name_set extra_opaque = ext.m_opaque;
|
||||
extra_opaque_pred pred([=](name const & n) { return extra_opaque.contains(n); });
|
||||
return std::unique_ptr<type_checker>(new type_checker(env, ngen, mk_default_converter(env,
|
||||
!ext.m_hide_module && relax_main_opaque,
|
||||
return std::unique_ptr<type_checker>(new type_checker(env, ngen, mk_default_converter(env, relax_main_opaque,
|
||||
true, pred)));
|
||||
}
|
||||
}
|
||||
}
|
||||
|
|
|
@ -22,11 +22,21 @@ namespace lean {
|
|||
environment hide_definition(environment const & env, name const & n);
|
||||
/** \brief Undo the modification made with \c hide_definition. */
|
||||
environment expose_definition(environment const & env, name const & n);
|
||||
/** \brief By default, the elaborator/unifier will treat the opaque definitions of the current/main
|
||||
module as transparent (when allowed). We can change this behavior using this function.
|
||||
/** \brief Return true iff \c n was exposed using \c expose_definition */
|
||||
bool is_exposed(environment const & env, name const & n);
|
||||
/** \brief Create a type checker that uses a converter based on the opaque/transparent hints
|
||||
provided by the user.
|
||||
|
||||
If \c relax_main_opaque is true, then all opaque definitions from the main module
|
||||
are treated as transparent.
|
||||
|
||||
If \c only_main_transparent is true, then only transparent definitions from the main module
|
||||
are treated as transparent.
|
||||
|
||||
The hints provided using #hide_definition and #expose_definition override this behavior.
|
||||
*/
|
||||
environment set_hide_main_opaque(environment const & env, bool f);
|
||||
bool get_hide_main_opaque(environment const & env);
|
||||
/** \brief Create a type checker that takes the hints into account. */
|
||||
std::unique_ptr<type_checker> mk_type_checker_with_hints(environment const & env, name_generator const & ngen, bool relax_main_opaque);
|
||||
std::unique_ptr<type_checker> mk_type_checker_with_hints(environment const & env,
|
||||
name_generator const & ngen,
|
||||
bool relax_main_opaque,
|
||||
bool only_main_transparent = false);
|
||||
}
|
||||
|
|
119
src/library/reducible.cpp
Normal file
119
src/library/reducible.cpp
Normal file
|
@ -0,0 +1,119 @@
|
|||
/*
|
||||
Copyright (c) 2014 Microsoft Corporation. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
|
||||
Author: Leonardo de Moura
|
||||
*/
|
||||
#include <string>
|
||||
#include "util/sstream.h"
|
||||
#include "kernel/environment.h"
|
||||
#include "kernel/type_checker.h"
|
||||
#include "library/kernel_serializer.h"
|
||||
#include "library/scoped_ext.h"
|
||||
#include "library/reducible.h"
|
||||
|
||||
namespace lean {
|
||||
|
||||
struct reducible_entry {
|
||||
reducible_status m_status;
|
||||
name m_name;
|
||||
reducible_entry():m_status(reducible_status::None) {}
|
||||
reducible_entry(reducible_status s, name const & n):m_status(s), m_name(n) {}
|
||||
};
|
||||
|
||||
struct reducible_state {
|
||||
name_set m_reducible_on;
|
||||
name_set m_reducible_off;
|
||||
|
||||
void add(reducible_entry const & e) {
|
||||
switch (e.m_status) {
|
||||
case reducible_status::On:
|
||||
m_reducible_on.insert(e.m_name);
|
||||
m_reducible_off.erase(e.m_name);
|
||||
break;
|
||||
case reducible_status::Off:
|
||||
m_reducible_on.erase(e.m_name);
|
||||
m_reducible_off.insert(e.m_name);
|
||||
break;
|
||||
case reducible_status::None:
|
||||
m_reducible_on.erase(e.m_name);
|
||||
m_reducible_off.erase(e.m_name);
|
||||
break;
|
||||
}
|
||||
}
|
||||
};
|
||||
|
||||
struct reducible_config {
|
||||
typedef reducible_state state;
|
||||
typedef reducible_entry entry;
|
||||
static void add_entry(environment const &, io_state const &, state & s, entry const & e) {
|
||||
s.add(e);
|
||||
}
|
||||
static name const & get_class_name() {
|
||||
static name g_class_name("reducible");
|
||||
return g_class_name;
|
||||
}
|
||||
static std::string const & get_serialization_key() {
|
||||
static std::string g_key("redu");
|
||||
return g_key;
|
||||
}
|
||||
static void write_entry(serializer & s, entry const & e) {
|
||||
s << static_cast<char>(e.m_status) << e.m_name;
|
||||
}
|
||||
static entry read_entry(deserializer & d) {
|
||||
entry e; char s;
|
||||
d >> s >> e.m_name;
|
||||
e.m_status = static_cast<reducible_status>(s);
|
||||
return e;
|
||||
}
|
||||
};
|
||||
|
||||
template class scoped_ext<reducible_config>;
|
||||
typedef scoped_ext<reducible_config> reducible_ext;
|
||||
|
||||
static void check_declaration(environment const & env, name const & n) {
|
||||
declaration const & d = env.get(n);
|
||||
if (!d.is_definition())
|
||||
throw exception(sstream() << "invalid reducible command, '" << n << "' is not a definition");
|
||||
if (d.is_theorem())
|
||||
throw exception(sstream() << "invalid reducible command, '" << n << "' is a theorem");
|
||||
if (d.is_opaque() && d.get_module_idx() != g_main_module_idx)
|
||||
throw exception(sstream() << "invalid reducible command, '" << n << "' is an opaque definition");
|
||||
}
|
||||
|
||||
environment set_reducible(environment const & env, name const & n, reducible_status s, bool persistent) {
|
||||
check_declaration(env, n);
|
||||
return reducible_ext::add_entry(env, get_dummy_ios(), reducible_entry(s, n), persistent);
|
||||
}
|
||||
|
||||
bool is_reducible_on(environment const & env, name const & n) {
|
||||
reducible_state const & s = reducible_ext::get_state(env);
|
||||
return s.m_reducible_on.contains(n);
|
||||
}
|
||||
|
||||
bool is_reducible_off(environment const & env, name const & n) {
|
||||
reducible_state const & s = reducible_ext::get_state(env);
|
||||
return s.m_reducible_off.contains(n);
|
||||
}
|
||||
|
||||
std::unique_ptr<type_checker> mk_type_checker(environment const & env, name_generator const & ngen,
|
||||
bool relax_main_opaque, bool only_main_reducible) {
|
||||
reducible_state const & s = reducible_ext::get_state(env);
|
||||
if (only_main_reducible) {
|
||||
name_set reducible_on = s.m_reducible_on;
|
||||
name_set reducible_off = s.m_reducible_off;
|
||||
extra_opaque_pred pred([=](name const & n) {
|
||||
return
|
||||
(!module::is_definition(env, n) || reducible_off.contains(n)) &&
|
||||
(!reducible_on.contains(n));
|
||||
});
|
||||
return std::unique_ptr<type_checker>(new type_checker(env, ngen, mk_default_converter(env, relax_main_opaque,
|
||||
true, pred)));
|
||||
} else {
|
||||
name_set reducible_off = s.m_reducible_off;
|
||||
extra_opaque_pred pred([=](name const & n) { return reducible_off.contains(n); });
|
||||
return std::unique_ptr<type_checker>(new type_checker(env, ngen, mk_default_converter(env, relax_main_opaque,
|
||||
true, pred)));
|
||||
}
|
||||
}
|
||||
}
|
39
src/library/reducible.h
Normal file
39
src/library/reducible.h
Normal file
|
@ -0,0 +1,39 @@
|
|||
/*
|
||||
Copyright (c) 2014 Microsoft Corporation. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
|
||||
Author: Leonardo de Moura
|
||||
*/
|
||||
#pragma once
|
||||
#include <memory>
|
||||
#include "kernel/type_checker.h"
|
||||
|
||||
namespace lean {
|
||||
enum class reducible_status { On, Off, None };
|
||||
/** \brief Mark the definition named \c n as reducible or not.
|
||||
|
||||
The method throws an exception if \c n is
|
||||
- not a definition
|
||||
- a theorem
|
||||
- an opaque definition that was not defined in main module
|
||||
|
||||
"Reducible" definitions can be freely unfolded by automation (i.e., elaborator, simplifier, etc).
|
||||
We should view it as a hint to automation.
|
||||
*/
|
||||
environment set_reducible(environment const & env, name const & n, reducible_status s, bool persistent = true);
|
||||
/** \brief Return true iff \c n was explicitly marked as reducible in the given environment.
|
||||
|
||||
\remark is_reducible_on(env, n) and is_reducible_off(env, n) cannot be simultaneously true,
|
||||
but both can be false (when set_reducible(env, n, None))
|
||||
*/
|
||||
bool is_reducible_on(environment const & env, name const & n);
|
||||
/** \brief Return true iff \c n was explicitly marked as not reducible in the given environment.
|
||||
|
||||
\see is_reducible_on
|
||||
*/
|
||||
bool is_reducible_off(environment const & env, name const & n);
|
||||
|
||||
/** \brief Create a type checker that takes the "reducibility" hints into account. */
|
||||
std::unique_ptr<type_checker> mk_type_checker(environment const & env, name_generator const & ngen,
|
||||
bool relax_main_opaque, bool only_main_reducible = false);
|
||||
}
|
|
@ -22,8 +22,9 @@ Author: Leonardo de Moura
|
|||
#include "kernel/error_msgs.h"
|
||||
#include "library/occurs.h"
|
||||
#include "library/locals.h"
|
||||
#include "library/module.h"
|
||||
#include "library/unifier.h"
|
||||
#include "library/opaque_hints.h"
|
||||
#include "library/reducible.h"
|
||||
#include "library/unifier_plugin.h"
|
||||
#include "library/kernel_bindings.h"
|
||||
#include "library/print.h"
|
||||
|
@ -47,7 +48,7 @@ unsigned get_unifier_max_steps(options const & opts) { return opts.get_unsigned(
|
|||
|
||||
static name g_unifier_computation {"unifier", "computation"};
|
||||
RegisterBoolOption(g_unifier_computation, LEAN_DEFAULT_UNIFIER_COMPUTATION,
|
||||
"(unifier) always case-split on reduction/computational steps when solving flex-rigid constraints");
|
||||
"(unifier) always case-split on reduction/computational steps when solving flex-rigid and delta-delta constraints");
|
||||
bool get_unifier_computation(options const & opts) { return opts.get_bool(g_unifier_computation, LEAN_DEFAULT_UNIFIER_COMPUTATION); }
|
||||
|
||||
static name g_unifier_expensive_classes {"unifier", "expensive_classes"};
|
||||
|
@ -295,6 +296,8 @@ struct unifier_fn {
|
|||
owned_map m_owned_map; // mapping from metavariable name m to delay factor of the choice constraint that owns m
|
||||
unifier_plugin m_plugin;
|
||||
type_checker_ptr m_tc[2];
|
||||
type_checker_ptr m_flex_rigid_tc[2]; // type checker used when solving flex rigid constraints. By default,
|
||||
// only the definitions from the main module are treated as transparent.
|
||||
unifier_config m_config;
|
||||
unsigned m_num_steps;
|
||||
bool m_pattern; //!< If true, then only higher-order (pattern) matching is used
|
||||
|
@ -391,8 +394,12 @@ struct unifier_fn {
|
|||
unifier_config const & cfg):
|
||||
m_env(env), m_ngen(ngen), m_subst(s), m_plugin(get_unifier_plugin(env)),
|
||||
m_config(cfg), m_num_steps(0), m_pattern(false) {
|
||||
m_tc[0] = mk_type_checker_with_hints(env, m_ngen.mk_child(), false);
|
||||
m_tc[1] = mk_type_checker_with_hints(env, m_ngen.mk_child(), true);
|
||||
m_tc[0] = mk_type_checker(env, m_ngen.mk_child(), false);
|
||||
m_tc[1] = mk_type_checker(env, m_ngen.mk_child(), true);
|
||||
if (!cfg.m_computation) {
|
||||
m_flex_rigid_tc[0] = mk_type_checker(env, m_ngen.mk_child(), false, true);
|
||||
m_flex_rigid_tc[1] = mk_type_checker(env, m_ngen.mk_child(), true, true);
|
||||
}
|
||||
m_next_assumption_idx = 0;
|
||||
m_next_cidx = 0;
|
||||
m_first = true;
|
||||
|
@ -559,6 +566,17 @@ struct unifier_fn {
|
|||
return r;
|
||||
}
|
||||
|
||||
expr flex_rigid_whnf(expr const & e, justification const & j, bool relax, buffer<constraint> & cs) {
|
||||
if (m_config.m_computation) {
|
||||
return whnf(e, j, relax, cs);
|
||||
} else {
|
||||
constraint_seq _cs;
|
||||
expr r = m_flex_rigid_tc[relax]->whnf(e, _cs);
|
||||
to_buffer(_cs, j, cs);
|
||||
return r;
|
||||
}
|
||||
}
|
||||
|
||||
justification mk_assign_justification(expr const & m, expr const & m_type, expr const & v_type, justification const & j) {
|
||||
auto r = j.get_main_expr();
|
||||
if (!r) r = m;
|
||||
|
@ -1178,18 +1196,20 @@ struct unifier_fn {
|
|||
}
|
||||
|
||||
justification a;
|
||||
bool relax = relax_main_opaque(c);
|
||||
if (m_config.m_computation || module::is_definition(m_env, d.get_name()) || is_reducible_on(m_env, d.get_name())) {
|
||||
// add case_split for t =?= s
|
||||
expr lhs_fn_val = instantiate_value_univ_params(d, const_levels(lhs_fn));
|
||||
expr rhs_fn_val = instantiate_value_univ_params(d, const_levels(rhs_fn));
|
||||
expr t = apply_beta(lhs_fn_val, lhs_args.size(), lhs_args.data());
|
||||
expr s = apply_beta(rhs_fn_val, rhs_args.size(), rhs_args.data());
|
||||
bool relax = relax_main_opaque(c);
|
||||
auto dcs = m_tc[relax]->is_def_eq(t, s, j);
|
||||
if (dcs.first) {
|
||||
// create a case split
|
||||
a = mk_assumption_justification(m_next_assumption_idx);
|
||||
add_case_split(std::unique_ptr<case_split>(new simple_case_split(*this, j, dcs.second.to_list())));
|
||||
}
|
||||
}
|
||||
|
||||
// process first case
|
||||
justification new_j = mk_composite1(j, a);
|
||||
|
@ -1245,6 +1265,18 @@ struct unifier_fn {
|
|||
buffer<constraints> & alts; // result: alternatives
|
||||
|
||||
optional<bool> _has_meta_args;
|
||||
|
||||
type_checker & tc() {
|
||||
return *u.m_tc[relax];
|
||||
}
|
||||
|
||||
type_checker & restricted_tc() {
|
||||
if (u.m_config.m_computation)
|
||||
return *u.m_tc[relax];
|
||||
else
|
||||
return *u.m_flex_rigid_tc[relax];
|
||||
}
|
||||
|
||||
/** \brief Return true if margs contains an expression \c e s.t. is_meta(e) */
|
||||
bool has_meta_args() {
|
||||
if (!_has_meta_args) {
|
||||
|
@ -1291,7 +1323,7 @@ struct unifier_fn {
|
|||
expr ensure_sufficient_args_core(expr mtype, unsigned i, constraint_seq & cs) {
|
||||
if (i == margs.size())
|
||||
return mtype;
|
||||
mtype = u.m_tc[relax]->ensure_pi(mtype, cs);
|
||||
mtype = tc().ensure_pi(mtype, cs);
|
||||
expr local = u.mk_local_for(mtype);
|
||||
expr body = instantiate(binding_body(mtype), local);
|
||||
return Pi(local, ensure_sufficient_args_core(body, i+1, cs));
|
||||
|
@ -1329,6 +1361,18 @@ struct unifier_fn {
|
|||
alts.push_back(cs.to_list());
|
||||
}
|
||||
|
||||
bool restricted_is_def_eq(expr const & marg, expr const & rhs, justification const & j, constraint_seq & cs) {
|
||||
try {
|
||||
if (restricted_tc().is_def_eq(marg, rhs, j, cs)) {
|
||||
return true;
|
||||
} else {
|
||||
return false;
|
||||
}
|
||||
} catch (exception & ex) {
|
||||
return false;
|
||||
}
|
||||
}
|
||||
|
||||
/**
|
||||
Given,
|
||||
m := a metavariable ?m
|
||||
|
@ -1351,8 +1395,8 @@ struct unifier_fn {
|
|||
auto new_mtype = ensure_sufficient_args(mtype, cs);
|
||||
// Remark: we should not use mk_eq_cnstr(marg, rhs, j) since is_def_eq may be able to reduce them.
|
||||
// The unifier assumes the eq constraints are reduced.
|
||||
if (u.m_tc[relax]->is_def_eq_types(marg, rhs, j, cs) &&
|
||||
u.m_tc[relax]->is_def_eq(marg, rhs, j, cs)) {
|
||||
if (tc().is_def_eq_types(marg, rhs, j, cs) &&
|
||||
restricted_is_def_eq(marg, rhs, j, cs)) {
|
||||
expr v = mk_lambda_for(new_mtype, mk_var(vidx));
|
||||
cs = cs + mk_eq_cnstr(m, v, j, relax);
|
||||
alts.push_back(cs.to_list());
|
||||
|
@ -1501,9 +1545,9 @@ struct unifier_fn {
|
|||
get_app_args(rhs, rargs);
|
||||
buffer<expr> sargs;
|
||||
try {
|
||||
expr f_type = u.m_tc[relax]->infer(f, cs);
|
||||
expr f_type = tc().infer(f, cs);
|
||||
for (expr const & rarg : rargs) {
|
||||
f_type = u.m_tc[relax]->ensure_pi(f_type, cs);
|
||||
f_type = tc().ensure_pi(f_type, cs);
|
||||
expr d_type = binding_domain(f_type);
|
||||
expr sarg = mk_imitiation_arg(rarg, d_type, locals, cs);
|
||||
sargs.push_back(sarg);
|
||||
|
@ -1578,13 +1622,13 @@ struct unifier_fn {
|
|||
try {
|
||||
// create a scope to make sure no constraints "leak" into the current state
|
||||
expr rhs_A = binding_domain(rhs);
|
||||
expr A_type = u.m_tc[relax]->infer(rhs_A, cs);
|
||||
expr A_type = tc().infer(rhs_A, cs);
|
||||
expr A = mk_imitiation_arg(rhs_A, A_type, locals, cs);
|
||||
expr local = mk_local(u.m_ngen.next(), binding_name(rhs), A, binding_info(rhs));
|
||||
locals.push_back(local);
|
||||
margs.push_back(local);
|
||||
expr rhs_B = instantiate(binding_body(rhs), local);
|
||||
expr B_type = u.m_tc[relax]->infer(rhs_B, cs);
|
||||
expr B_type = tc().infer(rhs_B, cs);
|
||||
expr B = mk_imitiation_arg(rhs_B, B_type, locals, cs);
|
||||
expr binding = is_pi(rhs) ? Pi(local, B) : Fun(local, B);
|
||||
locals.pop_back();
|
||||
|
@ -1725,7 +1769,7 @@ struct unifier_fn {
|
|||
process_flex_rigid_core(lhs, rhs, j, relax, alts);
|
||||
append_auxiliary_constraints(alts, to_list(aux.begin(), aux.end()));
|
||||
if (use_flex_rigid_whnf_split(lhs, rhs)) {
|
||||
expr rhs_whnf = whnf(rhs, j, relax, aux);
|
||||
expr rhs_whnf = flex_rigid_whnf(rhs, j, relax, aux);
|
||||
if (rhs_whnf != rhs) {
|
||||
if (is_meta(rhs_whnf)) {
|
||||
// it become a flex-flex constraint
|
||||
|
|
|
@ -73,20 +73,20 @@ variable dvd : Π (x y : nat) {H : not_zero y}, nat
|
|||
variables a b : nat
|
||||
|
||||
set_option pp.implicit true
|
||||
opaque add
|
||||
reducible add
|
||||
check dvd a (succ b)
|
||||
check (λ H : not_zero b, dvd a b)
|
||||
check (succ zero)
|
||||
check a + (succ zero)
|
||||
check dvd a (a + (succ b))
|
||||
|
||||
transparent add
|
||||
reducible [off] add
|
||||
check dvd a (a + (succ b))
|
||||
|
||||
opaque add
|
||||
reducible add
|
||||
check dvd a (a + (succ b))
|
||||
|
||||
transparent add
|
||||
reducible [off] add
|
||||
check dvd a (a + (succ b))
|
||||
|
||||
end nat
|
||||
|
|
|
@ -9,5 +9,5 @@ section
|
|||
definition concat (s t : list T) : list T
|
||||
:= list.rec t (fun x l u, list.cons x u) s
|
||||
|
||||
opaque concat
|
||||
reducible concat
|
||||
end
|
||||
|
|
|
@ -11,7 +11,7 @@ definition transport {A : Type} {a b : A} {P : A → Type} (p : a = b) (H : P a)
|
|||
theorem transport_refl {A : Type} {a : A} {P : A → Type} (H : P a) : transport (refl a) H = H
|
||||
:= refl H
|
||||
|
||||
opaque transport
|
||||
reducible [off] transport
|
||||
theorem transport_proof_irrel {A : Type} {a b : A} {P : A → Type} (p1 p2 : a = b) (H : P a) : transport p1 H = transport p2 H
|
||||
:= refl (transport p1 H)
|
||||
|
||||
|
|
|
@ -110,7 +110,7 @@ theorem length_concat (s t : list T) : length (s ++ t) = length s + length t :=
|
|||
list_induction_on s
|
||||
(calc
|
||||
length (concat nil t) = length t : refl _
|
||||
... = zero + length t : {symm add_zero_left}
|
||||
... = 0 + length t : {symm add_zero_left}
|
||||
... = length (@nil T) + length t : refl _)
|
||||
(take x s,
|
||||
assume H : length (concat s t) = length s + length t,
|
||||
|
|
Loading…
Reference in a new issue