feat(frontends/lean): split 'opaque_hint' command into 'opaque' and 'transparent'
Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
This commit is contained in:
parent
4b51d50ad4
commit
e3e2370a38
14 changed files with 44 additions and 61 deletions
|
@ -208,7 +208,7 @@ definition of_nat (n : ℕ) : ℤ := psub (pair n 0)
|
|||
theorem has_decidable_eq [instance] [protected] : decidable_eq ℤ :=
|
||||
take a b : ℤ, _
|
||||
|
||||
opaque_hint (hiding int)
|
||||
opaque int
|
||||
coercion of_nat
|
||||
|
||||
theorem eq_zero_intro (n : ℕ) : psub (pair n n) = 0 :=
|
||||
|
@ -302,7 +302,7 @@ have H5 : n + m = 0, from
|
|||
|
||||
---reverse equalities
|
||||
|
||||
opaque_hint (exposing int)
|
||||
transparent 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,
|
||||
|
@ -325,7 +325,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_hint (hiding int)
|
||||
opaque 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)) :
|
||||
|
@ -374,7 +374,7 @@ have H4 : n + m = 0, from
|
|||
add_eq_zero H4
|
||||
|
||||
--some of these had to be transparent for theorem cases
|
||||
opaque_hint (hiding psub proj)
|
||||
opaque psub proj
|
||||
|
||||
-- ## add
|
||||
|
||||
|
|
|
@ -72,7 +72,7 @@ theorem pred_zero : pred 0 = 0
|
|||
|
||||
theorem pred_succ {n : ℕ} : pred (succ n) = n
|
||||
|
||||
opaque_hint (hiding pred)
|
||||
opaque pred
|
||||
|
||||
theorem zero_or_succ_pred (n : ℕ) : n = 0 ∨ n = succ (pred n) :=
|
||||
induction_on n
|
||||
|
@ -159,7 +159,7 @@ theorem add_zero_right {n : ℕ} : n + 0 = n
|
|||
|
||||
theorem add_succ_right {n m : ℕ} : n + succ m = succ (n + m)
|
||||
|
||||
opaque_hint (hiding add)
|
||||
opaque add
|
||||
|
||||
theorem add_zero_left {n : ℕ} : 0 + n = n :=
|
||||
induction_on n
|
||||
|
@ -270,7 +270,7 @@ theorem mul_zero_right {n : ℕ} : n * 0 = 0
|
|||
|
||||
theorem mul_succ_right {n m : ℕ} : n * succ m = n * m + n
|
||||
|
||||
opaque_hint (hiding mul)
|
||||
opaque mul
|
||||
|
||||
-- ### commutativity, distributivity, associativity, identity
|
||||
|
||||
|
|
|
@ -30,7 +30,7 @@ exists_intro k H
|
|||
theorem le_elim {n m : ℕ} (H : n ≤ m) : ∃k, n + k = m :=
|
||||
H
|
||||
|
||||
opaque_hint (hiding le)
|
||||
opaque 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_hint (hiding sub)
|
||||
opaque sub
|
||||
|
||||
theorem sub_zero_left {n : ℕ} : 0 - n = 0 :=
|
||||
induction_on n sub_zero_right
|
||||
|
|
|
@ -193,8 +193,7 @@ 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_hint (hiding rec rec_constant rec_binary quotient_map quotient_map_binary)
|
||||
|
||||
opaque rec rec_constant rec_binary quotient_map quotient_map_binary
|
||||
|
||||
-- image
|
||||
-- -----
|
||||
|
@ -326,8 +325,4 @@ representative_map_to_quotient
|
|||
from subst (symm (and_absorb_left _ (reflR b))) H4,
|
||||
subst (symm (and_absorb_left _ (reflR a))) H5)
|
||||
|
||||
-- previously:
|
||||
-- opaque_hint (hiding fun_image rec is_quotient prelim_map)
|
||||
-- transparent: image, image_incl, quotient, quotient_abs, quotient_rep
|
||||
|
||||
end quotient
|
||||
|
|
|
@ -166,6 +166,4 @@ have Hy : pr2 (map_pair2 f (pair e e) v) = pr2 v, from
|
|||
... = pr2 v : Hid (pr2 v),
|
||||
prod.equal Hx Hy
|
||||
|
||||
opaque_hint (hiding flip map_pair map_pair2)
|
||||
|
||||
end quotient
|
||||
|
|
|
@ -48,7 +48,7 @@ eq_to_heq (eq.refl a)
|
|||
theorem heqt_elim {a : Prop} (H : a == true) : a :=
|
||||
eq_true_elim (heq_to_eq H)
|
||||
|
||||
opaque_hint (hiding cast)
|
||||
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 :=
|
||||
|
|
|
@ -8,7 +8,7 @@
|
|||
(require 'rx)
|
||||
|
||||
(defconst lean-keywords
|
||||
'("import" "abbreviation" "opaque_hint" "tactic_hint" "definition" "renaming"
|
||||
'("import" "abbreviation" "opaque" "transparent" "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"
|
||||
|
|
|
@ -44,7 +44,6 @@ static name g_hiding("hiding");
|
|||
static name g_exposing("exposing");
|
||||
static name g_renaming("renaming");
|
||||
static name g_as("as");
|
||||
static name g_module("[module]");
|
||||
static name g_colon(":");
|
||||
|
||||
environment print_cmd(parser & p) {
|
||||
|
@ -305,41 +304,30 @@ environment coercion_cmd(parser & p) {
|
|||
}
|
||||
}
|
||||
|
||||
environment opaque_hint_cmd(parser & p) {
|
||||
environment opaque_cmd_core(parser & p, bool hiding) {
|
||||
environment env = p.env();
|
||||
bool found = false;
|
||||
while (p.curr_is_token(g_lparen)) {
|
||||
p.next();
|
||||
bool hiding;
|
||||
auto pos = p.pos();
|
||||
if (p.curr_is_token_or_id(g_hiding))
|
||||
hiding = true;
|
||||
else if (p.curr_is_token_or_id(g_exposing))
|
||||
hiding = false;
|
||||
while (p.curr_is_identifier()) {
|
||||
name c = p.check_constant_next("invalid 'opaque/transparent' command, constant expected");
|
||||
found = true;
|
||||
if (hiding)
|
||||
env = hide_definition(env, c);
|
||||
else
|
||||
throw parser_error("invalid 'opaque_hint', 'hiding' or 'exposing' expected", pos);
|
||||
p.next();
|
||||
while (!p.curr_is_token(g_rparen)) {
|
||||
if (p.curr_is_token(g_module)) {
|
||||
found = true;
|
||||
p.next();
|
||||
env = set_hide_main_opaque(env, hiding);
|
||||
} else {
|
||||
name c = p.check_constant_next("invalid 'opaque_hint', constant expected");
|
||||
found = true;
|
||||
if (hiding)
|
||||
env = hide_definition(env, c);
|
||||
else
|
||||
env = expose_definition(env, c);
|
||||
}
|
||||
}
|
||||
p.next();
|
||||
env = expose_definition(env, c);
|
||||
}
|
||||
if (!found)
|
||||
throw exception("invalid empty 'opaque_hint' command");
|
||||
throw exception("invalid empty 'opaque/transparent' 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);
|
||||
|
@ -359,7 +347,8 @@ 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_hint", "add hints for the elaborator/unifier", opaque_hint_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("#erase_cache", "erase cached definition (for debugging purposes)", erase_cache_cmd));
|
||||
|
||||
register_decl_cmds(r);
|
||||
|
|
|
@ -79,8 +79,8 @@ token_table init_token_table() {
|
|||
|
||||
char const * commands[] = {"theorem", "axiom", "variable", "definition", "coercion",
|
||||
"variables", "[private]", "[protected]", "[inline]", "[visible]", "[instance]",
|
||||
"[class]", "[module]", "[coercion]",
|
||||
"abbreviation", "opaque_hint", "evaluate", "check", "print", "end", "namespace", "section", "import",
|
||||
"[class]", "[coercion]", "abbreviation", "opaque", "transparent",
|
||||
"evaluate", "check", "print", "end", "namespace", "section", "import",
|
||||
"abbreviation", "inductive", "record", "renaming", "extends", "structure", "module", "universe",
|
||||
"precedence", "infixl", "infixr", "infix", "postfix", "prefix", "notation", "context",
|
||||
"exit", "set_option", "open", "export", "calc_subst", "calc_refl", "calc_trans", "tactic_hint",
|
||||
|
|
|
@ -29,7 +29,7 @@ static environment update(environment const & env, opaque_hints_ext const & ext)
|
|||
static void check_definition(environment const & env, name const & n) {
|
||||
declaration d = env.get(n);
|
||||
if (!d.is_definition())
|
||||
throw exception(sstream() << "invalid opaque_hint, '" << n << "' is not a definition");
|
||||
throw exception(sstream() << "invalid opaque/transparent, '" << n << "' is not a definition");
|
||||
}
|
||||
environment hide_definition(environment const & env, name const & n) {
|
||||
check_definition(env, n);
|
||||
|
@ -41,7 +41,7 @@ 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_hint, '" << n << "' is not in the 'extra opaque' set");
|
||||
throw exception(sstream() << "invalid 'exposing' opaque/transparent, '" << n << "' is not in the 'extra opaque' set");
|
||||
ext.m_extra_opaque.erase(n);
|
||||
return update(env, ext);
|
||||
}
|
||||
|
@ -53,9 +53,11 @@ environment set_hide_main_opaque(environment const & env, bool f) {
|
|||
bool get_hide_main_opaque(environment const & env) {
|
||||
return get_extension(env).m_hide_module;
|
||||
}
|
||||
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) {
|
||||
auto const & ext = get_extension(env);
|
||||
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,
|
||||
!ext.m_hide_module && relax_main_opaque,
|
||||
true, ext.m_extra_opaque)));
|
||||
}
|
||||
}
|
||||
|
|
|
@ -73,20 +73,20 @@ variable dvd : Π (x y : nat) {H : not_zero y}, nat
|
|||
variables a b : nat
|
||||
|
||||
set_option pp.implicit true
|
||||
opaque_hint (hiding [module])
|
||||
opaque 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))
|
||||
|
||||
opaque_hint (exposing [module])
|
||||
transparent add
|
||||
check dvd a (a + (succ b))
|
||||
|
||||
opaque_hint (hiding add)
|
||||
opaque add
|
||||
check dvd a (a + (succ b))
|
||||
|
||||
opaque_hint (exposing add)
|
||||
transparent add
|
||||
check dvd a (a + (succ b))
|
||||
|
||||
end nat
|
||||
|
|
|
@ -9,6 +9,5 @@ section
|
|||
definition concat (s t : list T) : list T
|
||||
:= list.rec t (fun x l u, list.cons x u) s
|
||||
|
||||
opaque_hint (hiding concat)
|
||||
|
||||
opaque 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_hint (hiding transport)
|
||||
opaque 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)
|
||||
|
||||
|
|
Loading…
Reference in a new issue