feat(library/tactic): improve assumption tactic performance
This commit is contained in:
parent
393cefcf97
commit
7f0951b8e7
6 changed files with 147 additions and 174 deletions
|
@ -190,7 +190,7 @@ theorem s_mul_assoc_lemma_4 {n : ℕ+} {ε q : ℚ} (Hε : ε > 0) (Hq : q > 0)
|
||||||
let H3 := mul_le_of_le_div (pos_div_of_pos_of_pos Hq Hε) H2,
|
let H3 := mul_le_of_le_div (pos_div_of_pos_of_pos Hq Hε) H2,
|
||||||
rewrite -(one_mul ε),
|
rewrite -(one_mul ε),
|
||||||
apply mul_le_mul_of_mul_div_le,
|
apply mul_le_mul_of_mul_div_le,
|
||||||
exact H3
|
assumption
|
||||||
end
|
end
|
||||||
|
|
||||||
theorem of_nat_add (a b : ℕ) : of_nat (a + b) = of_nat a + of_nat b := sorry -- did Jeremy add this?
|
theorem of_nat_add (a b : ℕ) : of_nat (a + b) = of_nat a + of_nat b := sorry -- did Jeremy add this?
|
||||||
|
@ -327,10 +327,10 @@ theorem eq_of_bdd_var {s t : seq} (Hs : regular s) (Ht : regular t)
|
||||||
apply inv_pos
|
apply inv_pos
|
||||||
end
|
end
|
||||||
|
|
||||||
|
set_option pp.beta false
|
||||||
theorem pnat_bound {ε : ℚ} (Hε : ε > 0) : ∃ p : ℕ+, p⁻¹ ≤ ε :=
|
theorem pnat_bound {ε : ℚ} (Hε : ε > 0) : ∃ p : ℕ+, p⁻¹ ≤ ε :=
|
||||||
begin
|
begin
|
||||||
fapply exists.intro,
|
existsi (pceil (1 / ε)),
|
||||||
exact (pceil (1 / ε)),
|
|
||||||
rewrite -(rat.div_div (rat.ne_of_gt Hε)) at {2},
|
rewrite -(rat.div_div (rat.ne_of_gt Hε)) at {2},
|
||||||
apply pceil_helper,
|
apply pceil_helper,
|
||||||
apply pnat.le.refl
|
apply pnat.le.refl
|
||||||
|
@ -343,22 +343,20 @@ theorem bdd_of_eq_var {s t : seq} (Hs : regular s) (Ht : regular t) (Heq : s ≡
|
||||||
apply (exists.elim (pnat_bound Hε)),
|
apply (exists.elim (pnat_bound Hε)),
|
||||||
intro N HN,
|
intro N HN,
|
||||||
let Bd' := bdd_of_eq Heq N,
|
let Bd' := bdd_of_eq Heq N,
|
||||||
fapply exists.intro,
|
existsi 2 * N,
|
||||||
exact 2 * N,
|
|
||||||
intro n Hn,
|
intro n Hn,
|
||||||
apply rat.le.trans,
|
apply rat.le.trans,
|
||||||
apply Bd' n Hn,
|
apply Bd' n Hn,
|
||||||
apply HN
|
assumption
|
||||||
end
|
end
|
||||||
|
|
||||||
theorem equiv.trans (s t u : seq) (Hs : regular s) (Ht : regular t) (Hu : regular u)
|
theorem equiv.trans (s t u : seq) (Hs : regular s) (Ht : regular t) (Hu : regular u)
|
||||||
(H : s ≡ t) (H2 : t ≡ u) : s ≡ u :=
|
(H : s ≡ t) (H2 : t ≡ u) : s ≡ u :=
|
||||||
begin
|
begin
|
||||||
apply (eq_of_bdd Hs Hu),
|
apply eq_of_bdd Hs Hu,
|
||||||
intros,
|
intros,
|
||||||
fapply exists.intro,
|
existsi 2 * (2 * j),
|
||||||
exact 2 * (2 * j),
|
intro n Hn,
|
||||||
intros [n, Hn],
|
|
||||||
rewrite [-rat.sub_add_cancel (s n) (t n), rat.add.assoc],
|
rewrite [-rat.sub_add_cancel (s n) (t n), rat.add.assoc],
|
||||||
apply rat.le.trans,
|
apply rat.le.trans,
|
||||||
apply abs_add_le_abs_add_abs,
|
apply abs_add_le_abs_add_abs,
|
||||||
|
@ -366,8 +364,7 @@ theorem equiv.trans (s t u : seq) (Hs : regular s) (Ht : regular t) (Hu : regula
|
||||||
have Htu : abs (t n - u n) ≤ (2 * j)⁻¹, from bdd_of_eq H2 _ _ Hn,
|
have Htu : abs (t n - u n) ≤ (2 * j)⁻¹, from bdd_of_eq H2 _ _ Hn,
|
||||||
rewrite -(padd_halves j),
|
rewrite -(padd_halves j),
|
||||||
apply rat.add_le_add,
|
apply rat.add_le_add,
|
||||||
apply Hst, apply Htu
|
repeat assumption
|
||||||
-- assumption, assumption
|
|
||||||
end
|
end
|
||||||
|
|
||||||
-----------------------------------
|
-----------------------------------
|
||||||
|
@ -475,8 +472,8 @@ definition one : seq := λ n, 1
|
||||||
|
|
||||||
theorem s_add_comm (s t : seq) : sadd s t ≡ sadd t s :=
|
theorem s_add_comm (s t : seq) : sadd s t ≡ sadd t s :=
|
||||||
begin
|
begin
|
||||||
rewrite ↑sadd,
|
esimp [sadd],
|
||||||
intros n,
|
intro n,
|
||||||
rewrite [sub_add_eq_sub_sub, rat.add_sub_cancel, rat.sub_self, abs_zero],
|
rewrite [sub_add_eq_sub_sub, rat.add_sub_cancel, rat.sub_self, abs_zero],
|
||||||
apply add_invs_nonneg
|
apply add_invs_nonneg
|
||||||
end
|
end
|
||||||
|
@ -709,9 +706,8 @@ theorem add_well_defined {s t u v : seq} (Hs : regular s) (Ht : regular t) (Hu :
|
||||||
theorem mul_bound_helper {s t : seq} (Hs : regular s) (Ht : regular t) (a b c : ℕ+) (j : ℕ+) :
|
theorem mul_bound_helper {s t : seq} (Hs : regular s) (Ht : regular t) (a b c : ℕ+) (j : ℕ+) :
|
||||||
∃ N : ℕ+, ∀ n : ℕ+, n ≥ N → abs (s (a * n) * t (b * n) - s (c * n) * t (c * n)) ≤ j⁻¹ :=
|
∃ N : ℕ+, ∀ n : ℕ+, n ≥ N → abs (s (a * n) * t (b * n) - s (c * n) * t (c * n)) ≤ j⁻¹ :=
|
||||||
begin
|
begin
|
||||||
fapply exists.intro,
|
existsi pceil (((pnat.to_rat (K s)) * (b⁻¹ + c⁻¹) + (a⁻¹ + c⁻¹) *
|
||||||
exact (pceil (((pnat.to_rat (K s)) * (b⁻¹ + c⁻¹) + (a⁻¹ + c⁻¹) *
|
(pnat.to_rat (K t))) * (pnat.to_rat j)),
|
||||||
(pnat.to_rat (K t))) * (pnat.to_rat j))),
|
|
||||||
intros n Hn,
|
intros n Hn,
|
||||||
rewrite rewrite_helper4,
|
rewrite rewrite_helper4,
|
||||||
apply rat.le.trans,
|
apply rat.le.trans,
|
||||||
|
@ -776,32 +772,17 @@ theorem mul_bound_helper {s t : seq} (Hs : regular s) (Ht : regular t) (a b c :
|
||||||
theorem s_distrib {s t u : seq} (Hs : regular s) (Ht : regular t) (Hu : regular u) :
|
theorem s_distrib {s t u : seq} (Hs : regular s) (Ht : regular t) (Hu : regular u) :
|
||||||
smul s (sadd t u) ≡ sadd (smul s t) (smul s u) :=
|
smul s (sadd t u) ≡ sadd (smul s t) (smul s u) :=
|
||||||
begin
|
begin
|
||||||
/- apply eq_of_bdd,
|
|
||||||
apply reg_mul_reg,
|
|
||||||
eassumption,
|
|
||||||
apply reg_add_reg,
|
|
||||||
repeat eassumption,
|
|
||||||
apply reg_add_reg,
|
|
||||||
repeat eassumption,
|
|
||||||
apply reg_mul_reg,
|
|
||||||
repeat eassumption,
|
|
||||||
apply reg_mul_reg,
|
|
||||||
repeat eassumption,-/
|
|
||||||
apply eq_of_bdd,
|
apply eq_of_bdd,
|
||||||
apply reg_mul_reg,
|
apply reg_mul_reg,
|
||||||
apply Hs,
|
assumption,
|
||||||
apply reg_add_reg,
|
apply reg_add_reg,
|
||||||
apply Ht,
|
repeat assumption,
|
||||||
apply Hu,
|
|
||||||
apply reg_add_reg,
|
apply reg_add_reg,
|
||||||
|
repeat assumption,
|
||||||
apply reg_mul_reg,
|
apply reg_mul_reg,
|
||||||
rotate 2,
|
repeat assumption,
|
||||||
apply reg_mul_reg,
|
apply reg_mul_reg,
|
||||||
apply Hs,
|
repeat assumption,
|
||||||
apply Hu,
|
|
||||||
rotate 1,
|
|
||||||
apply Hs,
|
|
||||||
apply Ht,
|
|
||||||
intros,
|
intros,
|
||||||
let exh1 := λ a b c, mul_bound_helper Hs Ht a b c (2 * j),
|
let exh1 := λ a b c, mul_bound_helper Hs Ht a b c (2 * j),
|
||||||
apply exists.elim,
|
apply exists.elim,
|
||||||
|
@ -813,8 +794,7 @@ theorem s_distrib {s t u : seq} (Hs : regular s) (Ht : regular t) (Hu : regular
|
||||||
apply exh2,
|
apply exh2,
|
||||||
rotate 3,
|
rotate 3,
|
||||||
intros N2 HN2,
|
intros N2 HN2,
|
||||||
fapply exists.intro,
|
existsi max N1 N2,
|
||||||
exact (max N1 N2),
|
|
||||||
intros n Hn,
|
intros n Hn,
|
||||||
rewrite [↑sadd at *, ↑smul, rewrite_helper3, -padd_halves j, -*pnat_mul_assoc at *],
|
rewrite [↑sadd at *, ↑smul, rewrite_helper3, -padd_halves j, -*pnat_mul_assoc at *],
|
||||||
apply rat.le.trans,
|
apply rat.le.trans,
|
||||||
|
@ -841,9 +821,8 @@ theorem mul_zero_equiv_zero {s t : seq} (Hs : regular s) (Ht : regular t) (Htz :
|
||||||
(pos_div_of_pos_of_pos Hε (Kq_bound_pos Hs)),
|
(pos_div_of_pos_of_pos Hε (Kq_bound_pos Hs)),
|
||||||
apply exists.elim Bd,
|
apply exists.elim Bd,
|
||||||
intro N HN,
|
intro N HN,
|
||||||
fapply exists.intro,
|
existsi N,
|
||||||
exact N,
|
intro n Hn,
|
||||||
intros [n, Hn],
|
|
||||||
rewrite [↑equiv at Htz, ↑zero at *, rat.sub_zero, ↑smul, abs_mul],
|
rewrite [↑equiv at Htz, ↑zero at *, rat.sub_zero, ↑smul, abs_mul],
|
||||||
apply rat.le.trans,
|
apply rat.le.trans,
|
||||||
apply rat.mul_le_mul,
|
apply rat.mul_le_mul,
|
||||||
|
@ -882,8 +861,7 @@ theorem equiv_of_diff_equiv_zero {s t : seq} (Hs : regular s) (Ht : regular t)
|
||||||
apply eq_of_bdd Hs Ht,
|
apply eq_of_bdd Hs Ht,
|
||||||
intros,
|
intros,
|
||||||
let He := bdd_of_eq H,
|
let He := bdd_of_eq H,
|
||||||
fapply exists.intro,
|
existsi 2 * (2 * (2 * j)),
|
||||||
apply (2 * (2 * (2 * j))),
|
|
||||||
intros n Hn,
|
intros n Hn,
|
||||||
rewrite (rewrite_helper5 _ _ (s (2 * n)) (t (2 * n))),
|
rewrite (rewrite_helper5 _ _ (s (2 * n)) (t (2 * n))),
|
||||||
apply rat.le.trans,
|
apply rat.le.trans,
|
||||||
|
@ -921,16 +899,9 @@ theorem diff_equiv_zero_of_equiv {s t : seq} (Hs : regular s) (Ht : regular t) (
|
||||||
rotate 2,
|
rotate 2,
|
||||||
apply zero_is_reg,
|
apply zero_is_reg,
|
||||||
apply add_well_defined,
|
apply add_well_defined,
|
||||||
--repeat assumption,
|
repeat assumption,
|
||||||
apply Hs,
|
|
||||||
apply Hnt,
|
|
||||||
apply Ht,
|
|
||||||
apply Hnt,
|
|
||||||
apply H,
|
|
||||||
apply equiv.refl,
|
apply equiv.refl,
|
||||||
--repeat assumption
|
repeat assumption
|
||||||
apply Hsnt,
|
|
||||||
apply Htnt
|
|
||||||
end
|
end
|
||||||
|
|
||||||
theorem mul_well_defined_half1 {s t u : seq} (Hs : regular s) (Ht : regular t) (Hu : regular u)
|
theorem mul_well_defined_half1 {s t u : seq} (Hs : regular s) (Ht : regular t) (Hu : regular u)
|
||||||
|
@ -963,8 +934,7 @@ theorem mul_well_defined_half1 {s t u : seq} (Hs : regular s) (Ht : regular t) (
|
||||||
rotate 3,
|
rotate 3,
|
||||||
apply equiv.symm,
|
apply equiv.symm,
|
||||||
apply s_distrib,
|
apply s_distrib,
|
||||||
apply Hs, apply Ht, apply Hnu,
|
repeat assumption,
|
||||||
-- repeat assumption,
|
|
||||||
rotate 1,
|
rotate 1,
|
||||||
apply reg_add_reg Hst Hsnu,
|
apply reg_add_reg Hst Hsnu,
|
||||||
apply Hst,
|
apply Hst,
|
||||||
|
@ -976,12 +946,7 @@ theorem mul_well_defined_half1 {s t u : seq} (Hs : regular s) (Ht : regular t) (
|
||||||
apply mul_zero_equiv_zero,
|
apply mul_zero_equiv_zero,
|
||||||
rotate 2,
|
rotate 2,
|
||||||
apply diff_equiv_zero_of_equiv,
|
apply diff_equiv_zero_of_equiv,
|
||||||
apply Ht,
|
repeat assumption
|
||||||
apply Hu,
|
|
||||||
apply Etu,
|
|
||||||
apply Hs,
|
|
||||||
apply Htnu
|
|
||||||
-- repeat assumption
|
|
||||||
end
|
end
|
||||||
|
|
||||||
theorem mul_well_defined_half2 {s t u : seq} (Hs : regular s) (Ht : regular t) (Hu : regular u)
|
theorem mul_well_defined_half2 {s t u : seq} (Hs : regular s) (Ht : regular t) (Hu : regular u)
|
||||||
|
@ -1001,15 +966,7 @@ theorem mul_well_defined_half2 {s t u : seq} (Hs : regular s) (Ht : regular t) (
|
||||||
apply Ht,
|
apply Ht,
|
||||||
rotate 1,
|
rotate 1,
|
||||||
apply s_mul_comm,
|
apply s_mul_comm,
|
||||||
apply Hsu,
|
repeat assumption
|
||||||
apply Hus,
|
|
||||||
apply Htu,
|
|
||||||
apply Hus,
|
|
||||||
apply Hut,
|
|
||||||
apply Htu,
|
|
||||||
apply Hu,
|
|
||||||
apply Hs,
|
|
||||||
apply Est
|
|
||||||
end
|
end
|
||||||
|
|
||||||
theorem mul_well_defined {s t u v : seq} (Hs : regular s) (Ht : regular t) (Hu : regular u)
|
theorem mul_well_defined {s t u v : seq} (Hs : regular s) (Ht : regular t) (Hu : regular u)
|
||||||
|
@ -1020,11 +977,9 @@ theorem mul_well_defined {s t u v : seq} (Hs : regular s) (Ht : regular t) (Hu :
|
||||||
exact reg_mul_reg Hs Hv,
|
exact reg_mul_reg Hs Hv,
|
||||||
exact reg_mul_reg Hu Hv,
|
exact reg_mul_reg Hu Hv,
|
||||||
apply mul_well_defined_half1,
|
apply mul_well_defined_half1,
|
||||||
apply Hs, apply Ht, apply Hv, apply Etv,
|
repeat assumption,
|
||||||
-- repeat assumption,
|
|
||||||
apply mul_well_defined_half2,
|
apply mul_well_defined_half2,
|
||||||
-- repeat assumption
|
repeat assumption
|
||||||
apply Hs, apply Hu, apply Hv, apply Esu
|
|
||||||
end
|
end
|
||||||
|
|
||||||
theorem neg_well_defined {s t : seq} (Est : s ≡ t) : sneg s ≡ sneg t :=
|
theorem neg_well_defined {s t : seq} (Est : s ≡ t) : sneg s ≡ sneg t :=
|
||||||
|
|
|
@ -20,7 +20,7 @@ definition cons (a : A) (s : stream A) : stream A :=
|
||||||
|
|
||||||
notation h :: t := cons h t
|
notation h :: t := cons h t
|
||||||
|
|
||||||
definition head (s : stream A) : A :=
|
definition head [reducible] (s : stream A) : A :=
|
||||||
s 0
|
s 0
|
||||||
|
|
||||||
definition tail (s : stream A) : stream A :=
|
definition tail (s : stream A) : stream A :=
|
||||||
|
@ -29,7 +29,7 @@ definition tail (s : stream A) : stream A :=
|
||||||
definition drop (n : nat) (s : stream A) : stream A :=
|
definition drop (n : nat) (s : stream A) : stream A :=
|
||||||
λ i, s (i+n)
|
λ i, s (i+n)
|
||||||
|
|
||||||
definition nth (n : nat) (s : stream A) : A :=
|
definition nth [reducible] (n : nat) (s : stream A) : A :=
|
||||||
s n
|
s n
|
||||||
|
|
||||||
protected theorem eta (s : stream A) : head s :: tail s = s :=
|
protected theorem eta (s : stream A) : head s :: tail s = s :=
|
||||||
|
|
|
@ -120,11 +120,11 @@ static proof_state_seq apply_tactic_core(environment const & env, io_state const
|
||||||
auto mc = mk_class_instance_elaborator(
|
auto mc = mk_class_instance_elaborator(
|
||||||
env, ios, ctx, ngen.next(), optional<name>(),
|
env, ios, ctx, ngen.next(), optional<name>(),
|
||||||
use_local_insts, is_strict,
|
use_local_insts, is_strict,
|
||||||
some_expr(binding_domain(e_t)), e.get_tag(), cfg, nullptr);
|
some_expr(head_beta_reduce(binding_domain(e_t))), e.get_tag(), cfg, nullptr);
|
||||||
meta = mc.first;
|
meta = mc.first;
|
||||||
cs.push_back(mc.second);
|
cs.push_back(mc.second);
|
||||||
} else {
|
} else {
|
||||||
meta = g.mk_meta(ngen.next(), binding_domain(e_t));
|
meta = g.mk_meta(ngen.next(), head_beta_reduce(binding_domain(e_t)));
|
||||||
}
|
}
|
||||||
e = mk_app(e, meta);
|
e = mk_app(e, meta);
|
||||||
e_t = instantiate(binding_body(e_t), meta);
|
e_t = instantiate(binding_body(e_t), meta);
|
||||||
|
@ -197,33 +197,6 @@ tactic apply_tactic_core(expr const & e, constraint_seq const & cs) {
|
||||||
});
|
});
|
||||||
}
|
}
|
||||||
|
|
||||||
static tactic assumption_tactic_core(optional<unifier_kind> uk) {
|
|
||||||
return tactic([=](environment const & env, io_state const & ios, proof_state const & s) {
|
|
||||||
goals const & gs = s.get_goals();
|
|
||||||
if (empty(gs)) {
|
|
||||||
throw_no_goal_if_enabled(s);
|
|
||||||
return proof_state_seq();
|
|
||||||
}
|
|
||||||
proof_state new_s = s.update_report_failure(false);
|
|
||||||
proof_state_seq r;
|
|
||||||
goal g = head(gs);
|
|
||||||
buffer<expr> hs;
|
|
||||||
g.get_hyps(hs);
|
|
||||||
for (expr const & h : hs) {
|
|
||||||
r = append(r, apply_tactic_core(env, ios, new_s, h, DoNotAdd, IgnoreSubgoals, uk));
|
|
||||||
}
|
|
||||||
return r;
|
|
||||||
});
|
|
||||||
}
|
|
||||||
|
|
||||||
tactic eassumption_tactic() {
|
|
||||||
return assumption_tactic_core(optional<unifier_kind>());
|
|
||||||
}
|
|
||||||
|
|
||||||
tactic assumption_tactic() {
|
|
||||||
return assumption_tactic_core(optional<unifier_kind>(unifier_kind::Conservative));
|
|
||||||
}
|
|
||||||
|
|
||||||
tactic apply_tactic_core(elaborate_fn const & elab, expr const & e, add_meta_kind add_meta, subgoals_action_kind k) {
|
tactic apply_tactic_core(elaborate_fn const & elab, expr const & e, add_meta_kind add_meta, subgoals_action_kind k) {
|
||||||
return tactic([=](environment const & env, io_state const & ios, proof_state const & s) {
|
return tactic([=](environment const & env, io_state const & ios, proof_state const & s) {
|
||||||
goals const & gs = s.get_goals();
|
goals const & gs = s.get_goals();
|
||||||
|
@ -284,12 +257,6 @@ void initialize_apply_tactic() {
|
||||||
check_tactic_expr(app_arg(e), "invalid 'fapply' tactic, invalid argument");
|
check_tactic_expr(app_arg(e), "invalid 'fapply' tactic, invalid argument");
|
||||||
return fapply_tactic(fn, get_tactic_expr_expr(app_arg(e)));
|
return fapply_tactic(fn, get_tactic_expr_expr(app_arg(e)));
|
||||||
});
|
});
|
||||||
|
|
||||||
register_simple_tac(get_tactic_eassumption_name(),
|
|
||||||
[]() { return eassumption_tactic(); });
|
|
||||||
|
|
||||||
register_simple_tac(get_tactic_assumption_name(),
|
|
||||||
[]() { return assumption_tactic(); });
|
|
||||||
}
|
}
|
||||||
|
|
||||||
void finalize_apply_tactic() {
|
void finalize_apply_tactic() {
|
||||||
|
|
|
@ -32,7 +32,8 @@ bool solve_constraints(environment const & env, io_state const & ios, proof_stat
|
||||||
|
|
||||||
optional<expr> elaborate_with_respect_to(environment const & env, io_state const & ios, elaborate_fn const & elab,
|
optional<expr> elaborate_with_respect_to(environment const & env, io_state const & ios, elaborate_fn const & elab,
|
||||||
proof_state & s, expr const & e, optional<expr> const & _expected_type,
|
proof_state & s, expr const & e, optional<expr> const & _expected_type,
|
||||||
bool report_unassigned, bool enforce_type_during_elaboration) {
|
bool report_unassigned, bool enforce_type_during_elaboration,
|
||||||
|
bool conservative) {
|
||||||
name_generator ngen = s.get_ngen();
|
name_generator ngen = s.get_ngen();
|
||||||
substitution subst = s.get_subst();
|
substitution subst = s.get_subst();
|
||||||
goals const & gs = s.get_goals();
|
goals const & gs = s.get_goals();
|
||||||
|
@ -56,7 +57,7 @@ optional<expr> elaborate_with_respect_to(environment const & env, io_state const
|
||||||
} else {
|
} else {
|
||||||
to_buffer(s.get_postponed(), cs);
|
to_buffer(s.get_postponed(), cs);
|
||||||
if (expected_type) {
|
if (expected_type) {
|
||||||
auto tc = mk_type_checker(env, ngen.mk_child());
|
auto tc = mk_type_checker(env, ngen.mk_child(), conservative ? UnfoldReducible : UnfoldSemireducible);
|
||||||
auto e_t_cs = tc->infer(new_e);
|
auto e_t_cs = tc->infer(new_e);
|
||||||
expr t = *expected_type;
|
expr t = *expected_type;
|
||||||
e_t_cs.second.linearize(cs);
|
e_t_cs.second.linearize(cs);
|
||||||
|
|
|
@ -41,5 +41,6 @@ typedef std::function<elaborate_result(goal const &, name_generator &&, expr con
|
||||||
optional<expr> elaborate_with_respect_to(environment const & env, io_state const & ios, elaborate_fn const & elab,
|
optional<expr> elaborate_with_respect_to(environment const & env, io_state const & ios, elaborate_fn const & elab,
|
||||||
proof_state & s, expr const & e,
|
proof_state & s, expr const & e,
|
||||||
optional<expr> const & expected_type = optional<expr>(),
|
optional<expr> const & expected_type = optional<expr>(),
|
||||||
bool report_unassigned = false, bool enforce_type_during_elaboration = false);
|
bool report_unassigned = false, bool enforce_type_during_elaboration = false,
|
||||||
|
bool conservative = false);
|
||||||
}
|
}
|
||||||
|
|
|
@ -24,7 +24,8 @@ bool is_meta_placeholder(expr const & e) {
|
||||||
return std::all_of(args.begin(), args.end(), is_local);
|
return std::all_of(args.begin(), args.end(), is_local);
|
||||||
}
|
}
|
||||||
|
|
||||||
tactic exact_tactic(elaborate_fn const & elab, expr const & e, bool enforce_type_during_elaboration, bool allow_metavars) {
|
tactic exact_tactic(elaborate_fn const & elab, expr const & e, bool enforce_type_during_elaboration, bool allow_metavars,
|
||||||
|
bool conservative) {
|
||||||
return tactic01([=](environment const & env, io_state const & ios, proof_state const & s) {
|
return tactic01([=](environment const & env, io_state const & ios, proof_state const & s) {
|
||||||
proof_state new_s = s;
|
proof_state new_s = s;
|
||||||
goals const & gs = new_s.get_goals();
|
goals const & gs = new_s.get_goals();
|
||||||
|
@ -37,7 +38,8 @@ tactic exact_tactic(elaborate_fn const & elab, expr const & e, bool enforce_type
|
||||||
optional<expr> new_e;
|
optional<expr> new_e;
|
||||||
try {
|
try {
|
||||||
new_e = elaborate_with_respect_to(env, ios, elab, new_s, e, some_expr(t),
|
new_e = elaborate_with_respect_to(env, ios, elab, new_s, e, some_expr(t),
|
||||||
report_unassigned, enforce_type_during_elaboration);
|
report_unassigned, enforce_type_during_elaboration,
|
||||||
|
conservative);
|
||||||
} catch (exception &) {
|
} catch (exception &) {
|
||||||
if (s.report_failure())
|
if (s.report_failure())
|
||||||
throw;
|
throw;
|
||||||
|
@ -85,6 +87,48 @@ tactic exact_tactic(elaborate_fn const & elab, expr const & e, bool enforce_type
|
||||||
});
|
});
|
||||||
}
|
}
|
||||||
|
|
||||||
|
static tactic assumption_tactic_core(bool conservative) {
|
||||||
|
return tactic([=](environment const & env, io_state const & ios, proof_state const & s) {
|
||||||
|
goals const & gs = s.get_goals();
|
||||||
|
if (empty(gs)) {
|
||||||
|
throw_no_goal_if_enabled(s);
|
||||||
|
return proof_state_seq();
|
||||||
|
}
|
||||||
|
proof_state new_s = s.update_report_failure(false);
|
||||||
|
optional<tactic> tac;
|
||||||
|
goal g = head(gs);
|
||||||
|
buffer<expr> hs;
|
||||||
|
g.get_hyps(hs);
|
||||||
|
auto elab = [](goal const &, name_generator const &, expr const & H,
|
||||||
|
optional<expr> const &, substitution const & s, bool) -> elaborate_result {
|
||||||
|
return elaborate_result(H, s, constraints());
|
||||||
|
};
|
||||||
|
unsigned i = hs.size();
|
||||||
|
while (i > 0) {
|
||||||
|
--i;
|
||||||
|
expr const & h = hs[i];
|
||||||
|
tactic curr = exact_tactic(elab, h, false, false, conservative);
|
||||||
|
if (tac)
|
||||||
|
tac = orelse(*tac, curr);
|
||||||
|
else
|
||||||
|
tac = curr;
|
||||||
|
}
|
||||||
|
if (tac) {
|
||||||
|
return (*tac)(env, ios, s);
|
||||||
|
} else {
|
||||||
|
return proof_state_seq();
|
||||||
|
}
|
||||||
|
});
|
||||||
|
}
|
||||||
|
|
||||||
|
tactic eassumption_tactic() {
|
||||||
|
return assumption_tactic_core(false);
|
||||||
|
}
|
||||||
|
|
||||||
|
tactic assumption_tactic() {
|
||||||
|
return assumption_tactic_core(true);
|
||||||
|
}
|
||||||
|
|
||||||
static expr * g_exact_tac_fn = nullptr;
|
static expr * g_exact_tac_fn = nullptr;
|
||||||
static expr * g_rexact_tac_fn = nullptr;
|
static expr * g_rexact_tac_fn = nullptr;
|
||||||
static expr * g_refine_tac_fn = nullptr;
|
static expr * g_refine_tac_fn = nullptr;
|
||||||
|
@ -101,18 +145,23 @@ void initialize_exact_tactic() {
|
||||||
register_tac(exact_tac_name,
|
register_tac(exact_tac_name,
|
||||||
[](type_checker &, elaborate_fn const & fn, expr const & e, pos_info_provider const *) {
|
[](type_checker &, elaborate_fn const & fn, expr const & e, pos_info_provider const *) {
|
||||||
check_tactic_expr(app_arg(e), "invalid 'exact' tactic, invalid argument");
|
check_tactic_expr(app_arg(e), "invalid 'exact' tactic, invalid argument");
|
||||||
return exact_tactic(fn, get_tactic_expr_expr(app_arg(e)), true, false);
|
return exact_tactic(fn, get_tactic_expr_expr(app_arg(e)), true, false, false);
|
||||||
});
|
});
|
||||||
register_tac(rexact_tac_name,
|
register_tac(rexact_tac_name,
|
||||||
[](type_checker &, elaborate_fn const & fn, expr const & e, pos_info_provider const *) {
|
[](type_checker &, elaborate_fn const & fn, expr const & e, pos_info_provider const *) {
|
||||||
check_tactic_expr(app_arg(e), "invalid 'rexact' tactic, invalid argument");
|
check_tactic_expr(app_arg(e), "invalid 'rexact' tactic, invalid argument");
|
||||||
return exact_tactic(fn, get_tactic_expr_expr(app_arg(e)), false, false);
|
return exact_tactic(fn, get_tactic_expr_expr(app_arg(e)), false, false, false);
|
||||||
});
|
});
|
||||||
register_tac(refine_tac_name,
|
register_tac(refine_tac_name,
|
||||||
[](type_checker &, elaborate_fn const & fn, expr const & e, pos_info_provider const *) {
|
[](type_checker &, elaborate_fn const & fn, expr const & e, pos_info_provider const *) {
|
||||||
check_tactic_expr(app_arg(e), "invalid 'refine' tactic, invalid argument");
|
check_tactic_expr(app_arg(e), "invalid 'refine' tactic, invalid argument");
|
||||||
return exact_tactic(fn, get_tactic_expr_expr(app_arg(e)), true, true);
|
return exact_tactic(fn, get_tactic_expr_expr(app_arg(e)), true, true, false);
|
||||||
});
|
});
|
||||||
|
register_simple_tac(get_tactic_eassumption_name(),
|
||||||
|
[]() { return eassumption_tactic(); });
|
||||||
|
|
||||||
|
register_simple_tac(get_tactic_assumption_name(),
|
||||||
|
[]() { return assumption_tactic(); });
|
||||||
}
|
}
|
||||||
void finalize_exact_tactic() {
|
void finalize_exact_tactic() {
|
||||||
delete g_exact_tac_fn;
|
delete g_exact_tac_fn;
|
||||||
|
|
Loading…
Reference in a new issue