feat(frontends/lean): add simp tactic frontend stub

This commit also removes the fake_simplifier. It doesn't work anymore
because simp is now a reserved word.
This commit is contained in:
Leonardo de Moura 2015-07-14 09:54:53 -04:00
parent 5034de9c4e
commit 3ab0e07ba9
21 changed files with 94 additions and 37 deletions

View file

@ -27,10 +27,9 @@ following:
-/
import data.nat.basic data.nat.order data.nat.sub data.prod
import algebra.relation algebra.binary algebra.ordered_ring
import tools.fake_simplifier
open eq.ops
open prod relation nat
open decidable binary fake_simplifier
open decidable binary
/- the type of integers -/

View file

@ -9,7 +9,6 @@ and transfer the results.
import .basic algebra.ordered_ring
open nat
open decidable
open fake_simplifier
open int eq.ops
namespace int

View file

@ -5,8 +5,8 @@ Authors: Jeremy Avigad, Leonardo de Moura
Definitions and properties of div and mod. Much of the development follows Isabelle's library.
-/
import data.nat.sub tools.fake_simplifier
open eq.ops well_founded decidable fake_simplifier prod
import data.nat.sub
open eq.ops well_founded decidable prod
namespace nat

View file

@ -6,7 +6,7 @@ Authors: Jeremy Avigad, Leonardo de Moura
Definitions and properties of gcd, lcm, and coprime.
-/
import .div
open eq.ops well_founded decidable fake_simplifier prod
open eq.ops well_founded decidable prod
namespace nat

View file

@ -6,9 +6,7 @@ Authors: Floris van Doorn, Jeremy Avigad
Subtraction on the natural numbers, as well as min, max, and distance.
-/
import .order
import tools.fake_simplifier
open eq.ops
open fake_simplifier
namespace nat

View file

@ -219,9 +219,9 @@ theorem add_halves (p : +) : (2 * p)⁻¹ + (2 * p)⁻¹ = p⁻¹ :=
theorem add_halves_double (m n : +) :
m⁻¹ + n⁻¹ = ((2 * m)⁻¹ + (2 * n)⁻¹) + ((2 * m)⁻¹ + (2 * n)⁻¹) :=
have simp [visible] : ∀ a b : , (a + a) + (b + b) = (a + b) + (a + b),
have hsimp [visible] : ∀ a b : , (a + a) + (b + b) = (a + b) + (a + b),
by intros; rewrite [rat.add.assoc, -(rat.add.assoc a b b), {_+b}rat.add.comm, -*rat.add.assoc],
by rewrite [-add_halves m, -add_halves n, simp]
by rewrite [-add_halves m, -add_halves n, hsimp]
theorem inv_mul_eq_mul_inv {p q : +} : (p * q)⁻¹ = p⁻¹ * q⁻¹ :=
by rewrite [↑inv, pnat_to_rat_mul, one_div_mul_one_div''']
@ -297,8 +297,8 @@ theorem div_le_pnat (q : ) (n : +) (H : q ≥ n⁻¹) : 1 / q ≤ rat_of_p
theorem pnat_cancel' (n m : +) : (n * n * m)⁻¹ * (rat_of_pnat n * rat_of_pnat n) = m⁻¹ :=
begin
have simp : ∀ a b c : , (a * a * (b * b * c)) = (a * b) * (a * b) * c, from sorry, -- simp
rewrite [rat.mul.comm, *inv_mul_eq_mul_inv, simp, *inv_cancel_left, *rat.one_mul]
have hsimp : ∀ a b c : , (a * a * (b * b * c)) = (a * b) * (a * b) * c, from sorry, -- simp
rewrite [rat.mul.comm, *inv_mul_eq_mul_inv, hsimp, *inv_cancel_left, *rat.one_mul]
end
definition pceil (a : ) : + := pnat.pos (ubound a) !ubound_pos

View file

@ -4,10 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Author: Floris van Doorn
-/
import logic ..prod algebra.relation
import tools.fake_simplifier
open prod eq.ops
open fake_simplifier
namespace quotient

View file

@ -324,10 +324,10 @@ section migrate_algebra
end migrate_algebra
theorem rat_of_nat_abs (a : ) : abs (of_int a) = of_nat (int.nat_abs a) :=
have simp [visible] : ∀ n : , of_int (int.neg_succ_of_nat n) = - of_nat (nat.succ n), from λ n, rfl,
have hsimp [visible] : ∀ n : , of_int (int.neg_succ_of_nat n) = - of_nat (nat.succ n), from λ n, rfl,
int.induction_on a
(take b, abs_of_nonneg (!of_nat_nonneg))
(take b, by rewrite [simp, abs_neg, abs_of_nonneg (!of_nat_nonneg)])
(take b, by rewrite [hsimp, abs_neg, abs_of_nonneg (!of_nat_nonneg)])
definition ubound : := λ a : , nat.succ (int.nat_abs (num a))

View file

@ -164,9 +164,9 @@ theorem eq_of_bdd {s t : seq} (Hs : regular s) (Ht : regular t)
apply rat.add_le_add,
apply HNj (max j Nj) (max_right j Nj),
apply Ht,
have simp : ∀ m : +, n⁻¹ + m⁻¹ + (j⁻¹ + (m⁻¹ + n⁻¹)) = (n⁻¹ + n⁻¹ + j⁻¹) + (m⁻¹ + m⁻¹),
have hsimp : ∀ m : +, n⁻¹ + m⁻¹ + (j⁻¹ + (m⁻¹ + n⁻¹)) = (n⁻¹ + n⁻¹ + j⁻¹) + (m⁻¹ + m⁻¹),
from sorry, -- simplifier
rewrite simp,
rewrite hsimp,
have Hms : (max j Nj)⁻¹ + (max j Nj)⁻¹ ≤ j⁻¹ + j⁻¹, begin
apply rat.add_le_add,
apply inv_ge_of_le (max_left j Nj),
@ -733,7 +733,7 @@ theorem mul_neg_equiv_neg_mul {s t : seq} : smul s (sneg t) ≡ sneg (smul s t)
theorem equiv_of_diff_equiv_zero {s t : seq} (Hs : regular s) (Ht : regular t)
(H : sadd s (sneg t) ≡ zero) : s ≡ t :=
begin
have simp : ∀ a b c d e : , a + b + c + (d + e) = (b + d) + a + e + c, from sorry,
have hsimp : ∀ a b c d e : , a + b + c + (d + e) = (b + d) + a + e + c, from sorry,
apply eq_of_bdd Hs Ht,
intros,
let He := bdd_of_eq H,
@ -749,7 +749,7 @@ theorem equiv_of_diff_equiv_zero {s t : seq} (Hs : regular s) (Ht : regular t)
let He' := λ a b c, !rat.sub_zero ▸ (He a b c),
apply (He' _ _ Hn),
apply Ht,
rewrite [simp, add_halves, -(add_halves j), -(add_halves (2 * j)), -*rat.add.assoc],
rewrite [hsimp, add_halves, -(add_halves j), -(add_halves (2 * j)), -*rat.add.assoc],
apply rat.add_le_add_right,
apply add_le_add_three,
repeat (apply rat.le.trans; apply inv_ge_of_le Hn; apply inv_two_mul_le_inv)

View file

@ -99,7 +99,7 @@ definition krewrite_tac (e : expr_list) : tactic := builtin
-- - x : exclude the give global rewrites
-- - t : tactic for discharging conditions
-- - l : location
definition simp_tac (d a : bool) (e : expr_list) (n : identifier_list) (x : identifier_list) (t : option tactic) (l : expr) : tactic := builtin
definition simp_tac (e : expr_list) (n : identifier_list) (x : identifier_list) (t : option tactic) (l : expr) : tactic := builtin
-- with_options_tac is just a marker for the builtin 'with_options' notation
definition with_options_tac (o : expr) (t : tactic) : tactic := builtin

View file

@ -1,8 +0,0 @@
open tactic
namespace fake_simplifier
-- until we have the simplifier...
definition simp : tactic := apply sorry
end fake_simplifier

View file

@ -9,6 +9,7 @@ coercion_elaborator.cpp info_tactic.cpp
init_module.cpp elaborator_context.cpp calc_proof_elaborator.cpp
parse_tactic_location.cpp parse_rewrite_tactic.cpp builtin_tactics.cpp
type_util.cpp elaborator_exception.cpp migrate_cmd.cpp local_ref_info.cpp
obtain_expr.cpp decl_attributes.cpp nested_declaration.cpp parse_with_options_tactic.cpp)
obtain_expr.cpp decl_attributes.cpp nested_declaration.cpp
parse_with_options_tactic.cpp parse_simp_tactic.cpp)
target_link_libraries(lean_frontend ${LEAN_LIBS})

View file

@ -10,6 +10,7 @@ Author: Leonardo de Moura
#include "frontends/lean/parser.h"
#include "frontends/lean/parse_rewrite_tactic.h"
#include "frontends/lean/parse_with_options_tactic.h"
#include "frontends/lean/parse_simp_tactic.h"
namespace lean {
namespace notation {
@ -52,6 +53,10 @@ static expr parse_with_options_tactic_expr(parser & p, unsigned, expr const *, p
return p.save_pos(parse_with_options_tactic(p), pos);
}
static expr parse_simp_tactic_expr(parser & p, unsigned, expr const *, pos_info const & pos) {
return p.save_pos(parse_simp_tactic(p), pos);
}
static expr parse_generalize_tactic(parser & p, unsigned, expr const *, pos_info const & pos) {
expr e = p.parse_tactic_expr_arg();
name id;
@ -83,6 +88,7 @@ parse_table init_tactic_nud_table() {
r = r.add({transition("fold", mk_ext_action(parse_fold_tactic_expr))}, x0);
r = r.add({transition("let", mk_ext_action(parse_let_tactic))}, x0);
r = r.add({transition("with_options", mk_ext_action(parse_with_options_tactic_expr))}, x0);
r = r.add({transition("simp", mk_ext_action(parse_simp_tactic_expr))}, x0);
return r;
}

View file

@ -0,0 +1,17 @@
/*
Copyright (c) 2015 Microsoft Corporation. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Author: Leonardo de Moura
*/
#include "library/tactic/exact_tactic.h"
#include "frontends/lean/parser.h"
namespace lean {
expr parse_simp_tactic(parser & p) {
// TODO(Leo)
auto pos = p.pos();
expr sorry = p.mk_sorry(pos);
return p.mk_app(get_exact_tac_fn(), sorry, pos);
}
}

View file

@ -0,0 +1,13 @@
/*
Copyright (c) 2015 Microsoft Corporation. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Author: Leonardo de Moura
*/
#pragma once
#include "kernel/expr.h"
namespace lean {
class parser;
expr parse_simp_tactic(parser & p);
}

View file

@ -95,7 +95,7 @@ void init_token_table(token_table & t) {
{"{|", g_max_prec}, {"|}", 0}, {"", 0}, {"", g_max_prec}, {"", 0}, {"^", 0}, {"", 0}, {"", 0},
{"using", 0}, {"|", 0}, {"!", g_max_prec}, {"?", 0}, {"with", 0}, {"...", 0}, {",", 0},
{".", 0}, {":", 0}, {"::", 0}, {"calc", 0}, {"rewrite", 0}, {"xrewrite", 0}, {"krewrite", 0},
{"esimp", 0}, {"fold", 0}, {"unfold", 0}, {"with_options", 0},
{"esimp", 0}, {"fold", 0}, {"unfold", 0}, {"with_options", 0}, {"simp", 0},
{"generalize", 0}, {"as", 0}, {":=", 0}, {"--", 0}, {"#", 0},
{"(*", 0}, {"/-", 0}, {"begin", g_max_prec}, {"begin+", g_max_prec}, {"abstract", g_max_prec},
{"proof", g_max_prec}, {"qed", 0}, {"@", g_max_prec},

View file

@ -1,2 +1,2 @@
add_library(simplifier ceqv.cpp rewrite_rule_set.cpp init_module.cpp)
add_library(simplifier ceqv.cpp rewrite_rule_set.cpp init_module.cpp simp_tactic.cpp)
target_link_libraries(simplifier ${LEAN_LIBS})

View file

@ -0,0 +1,15 @@
/*
Copyright (c) 2015 Microsoft Corporation. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Author: Leonardo de Moura
*/
#include "library/simplifier/simp_tactic.h"
namespace lean {
expr mk_simp_tactic_expr(buffer<expr> const & /* ls */, buffer<name> const & /* ns */,
buffer<name> const & /* ex */, optional<expr> const & /* pre_tac */,
location const & /* loc */) {
return expr();
}
}

View file

@ -0,0 +1,21 @@
/*
Copyright (c) 2015 Microsoft Corporation. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Author: Leonardo de Moura
*/
#pragma once
#include "kernel/expr.h"
#include "library/tactic/location.h"
namespace lean {
/** \brief Create a simp tactic expression where
- ls is a collection of pre-terms representing additional lemmas that should be used as rewriting rules.
- ns is a collection of namespaces that should provide rewriting rules for this tactic.
- ex is a collection of global rewriting rules that should be excluded.
- pre_tac (if provided) is a tactic used to discharge assumptions in conditional rewriting rules.
- loc is the scope of the tactic (i.e., which hypothesis and/or conclusion will be simplified).
*/
expr mk_simp_tactic_expr(buffer<expr> const & ls, buffer<name> const & ns, buffer<name> const & ex,
optional<expr> const & pre_tac, location const & loc);
}

View file

@ -1,8 +1,7 @@
import logic data.nat.sub algebra.relation data.prod
import tools.fake_simplifier
open nat relation relation.iff_ops prod
open fake_simplifier decidable
open decidable
open eq.ops
namespace nat

View file

@ -7,7 +7,7 @@
import logic data.nat
open nat
namespace simp
namespace nsimp
-- set_option pp.universes true
-- set_option pp.implicit true
@ -25,4 +25,4 @@ theorem simp_app [instance] (S T : Type) (f1 f2 : S → T) (s1 s2 : S)
(C1 : simplifies_to f1 f2) (C2 : simplifies_to s1 s2) : simplifies_to (f1 s1) (f2 s2) :=
simplifies_to.mk (congr (get_eq C1) (get_eq C2))
end simp
end nsimp