refactor(builtin): remove dead module heq
Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
This commit is contained in:
parent
0283887ee9
commit
f4ec874c6e
27 changed files with 50 additions and 71 deletions
|
@ -65,7 +65,7 @@ theorem dep_if_elim_else {A : TypeU} (c : Bool) (t : c → A) (e : ¬ c → A) (
|
|||
theorem dep_if_false {A : TypeU} (t : false → A) (e : ¬ false → A) : dep_if false t e = e not_false_trivial
|
||||
:= dep_if_elim_else false t e not_false_trivial
|
||||
|
||||
import heq
|
||||
set_option simplifier::heq true -- enable heterogeneous equality support
|
||||
|
||||
theorem dep_if_congr {A : TypeM} (c1 c2 : Bool)
|
||||
(t1 : c1 → A) (t2 : c2 → A)
|
||||
|
|
|
@ -93,10 +93,8 @@ add_kernel_theory("Nat.lean" ${CMAKE_CURRENT_BINARY_DIR}/kernel.olean)
|
|||
add_theory("Int.lean" "${CMAKE_CURRENT_BINARY_DIR}/Nat.olean")
|
||||
add_theory("Real.lean" "${CMAKE_CURRENT_BINARY_DIR}/Int.olean")
|
||||
add_theory("specialfn.lean" "${CMAKE_CURRENT_BINARY_DIR}/Real.olean")
|
||||
add_theory("heq.lean" "${CMAKE_CURRENT_BINARY_DIR}/Nat.olean")
|
||||
|
||||
update_interface("kernel.olean" "kernel" "-n")
|
||||
update_interface("Nat.olean" "library/arith" "-n")
|
||||
update_interface("Int.olean" "library/arith" "")
|
||||
update_interface("Real.olean" "library/arith" "")
|
||||
update_interface("heq.olean" "library" "")
|
||||
|
|
|
@ -1,4 +0,0 @@
|
|||
-- Axioms and theorems for casting and heterogenous equality
|
||||
import macros
|
||||
|
||||
|
Binary file not shown.
|
@ -1,6 +1,6 @@
|
|||
add_library(library kernel_bindings.cpp deep_copy.cpp
|
||||
context_to_lambda.cpp placeholder.cpp expr_lt.cpp substitution.cpp
|
||||
fo_unify.cpp bin_op.cpp equality.cpp io_state_stream.cpp printer.cpp
|
||||
hop_match.cpp heq_decls.cpp)
|
||||
hop_match.cpp)
|
||||
|
||||
target_link_libraries(library ${LEAN_LIBS})
|
||||
|
|
|
@ -1,9 +0,0 @@
|
|||
/*
|
||||
Copyright (c) 2013 Microsoft Corporation. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
*/
|
||||
// Automatically generated file, DO NOT EDIT
|
||||
#include "kernel/environment.h"
|
||||
#include "kernel/decl_macros.h"
|
||||
namespace lean {
|
||||
}
|
|
@ -1,8 +0,0 @@
|
|||
/*
|
||||
Copyright (c) 2013 Microsoft Corporation. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
*/
|
||||
// Automatically generated file, DO NOT EDIT
|
||||
#include "kernel/expr.h"
|
||||
namespace lean {
|
||||
}
|
|
@ -19,7 +19,6 @@ Author: Leonardo de Moura
|
|||
#include "kernel/kernel.h"
|
||||
#include "kernel/max_sharing.h"
|
||||
#include "kernel/occurs.h"
|
||||
#include "library/heq_decls.h"
|
||||
#include "library/kernel_bindings.h"
|
||||
#include "library/expr_pair.h"
|
||||
#include "library/hop_match.h"
|
||||
|
@ -63,6 +62,10 @@ Author: Leonardo de Moura
|
|||
#define LEAN_SIMPLIFIER_MEMOIZE true
|
||||
#endif
|
||||
|
||||
#ifndef LEAN_SIMPLIFIER_HEQ
|
||||
#define LEAN_SIMPLIFIER_HEQ false
|
||||
#endif
|
||||
|
||||
#ifndef LEAN_SIMPLIFIER_PRESERVE_BINDER_NAMES
|
||||
#define LEAN_SIMPLIFIER_PRESERVE_BINDER_NAMES true
|
||||
#endif
|
||||
|
@ -82,6 +85,7 @@ static name g_simplifier_unfold {"simplifier", "unfold"};
|
|||
static name g_simplifier_conditional {"simplifier", "conditional"};
|
||||
static name g_simplifier_memoize {"simplifier", "memoize"};
|
||||
static name g_simplifier_max_steps {"simplifier", "max_steps"};
|
||||
static name g_simplifier_heq {"simplifier", "heq"};
|
||||
static name g_simplifier_preserve_binder_names {"simplifier", "preserve_binder_names"};
|
||||
|
||||
RegisterBoolOption(g_simplifier_proofs, LEAN_SIMPLIFIER_PROOFS, "(simplifier) generate proofs");
|
||||
|
@ -95,6 +99,7 @@ RegisterBoolOption(g_simplifier_conditional, LEAN_SIMPLIFIER_CONDITIONAL, "(simp
|
|||
RegisterBoolOption(g_simplifier_memoize, LEAN_SIMPLIFIER_MEMOIZE, "(simplifier) memoize/cache intermediate results");
|
||||
RegisterBoolOption(g_simplifier_preserve_binder_names, LEAN_SIMPLIFIER_PRESERVE_BINDER_NAMES,
|
||||
"(simplifier) (try to) preserve binder names when applying higher-order rewrite rules");
|
||||
RegisterBoolOption(g_simplifier_heq, LEAN_SIMPLIFIER_HEQ, "(simplifier) use heterogeneous equality support");
|
||||
RegisterUnsignedOption(g_simplifier_max_steps, LEAN_SIMPLIFIER_MAX_STEPS, "(simplifier) maximum number of steps");
|
||||
|
||||
bool get_simplifier_proofs(options const & opts) { return opts.get_bool(g_simplifier_proofs, LEAN_SIMPLIFIER_PROOFS); }
|
||||
|
@ -106,6 +111,7 @@ bool get_simplifier_eval(options const & opts) { return opts.get_bool(g_simplifi
|
|||
bool get_simplifier_unfold(options const & opts) { return opts.get_bool(g_simplifier_unfold, LEAN_SIMPLIFIER_UNFOLD); }
|
||||
bool get_simplifier_conditional(options const & opts) { return opts.get_bool(g_simplifier_conditional, LEAN_SIMPLIFIER_CONDITIONAL); }
|
||||
bool get_simplifier_memoize(options const & opts) { return opts.get_bool(g_simplifier_memoize, LEAN_SIMPLIFIER_MEMOIZE); }
|
||||
bool get_simplifier_heq(options const & opts) { return opts.get_bool(g_simplifier_heq, LEAN_SIMPLIFIER_HEQ); }
|
||||
bool get_simplifier_preserve_binder_names(options const & opts) {
|
||||
return opts.get_bool(g_simplifier_preserve_binder_names, LEAN_SIMPLIFIER_PRESERVE_BINDER_NAMES);
|
||||
}
|
||||
|
@ -129,7 +135,6 @@ class simplifier_cell::imp {
|
|||
ro_environment m_env;
|
||||
options m_options;
|
||||
type_checker m_tc;
|
||||
bool m_has_heq;
|
||||
rule_sets m_rule_sets;
|
||||
cache m_cache;
|
||||
max_sharing_fn m_max_sharing;
|
||||
|
@ -153,6 +158,7 @@ class simplifier_cell::imp {
|
|||
bool m_conditional;
|
||||
bool m_memoize;
|
||||
bool m_preserve_binder_names;
|
||||
bool m_use_heq;
|
||||
unsigned m_max_steps;
|
||||
|
||||
struct updt_rule_set {
|
||||
|
@ -458,7 +464,7 @@ class simplifier_cell::imp {
|
|||
}
|
||||
|
||||
result simplify_app(expr const & e) {
|
||||
if (m_has_heq && is_cast(e)) {
|
||||
if (m_use_heq && is_cast(e)) {
|
||||
// e is of the form (cast A B H a)
|
||||
// a : A
|
||||
// e : B
|
||||
|
@ -628,7 +634,7 @@ class simplifier_cell::imp {
|
|||
buffer<expr> new_args;
|
||||
buffer<optional<expr>> proofs; // used only if m_proofs_enabled
|
||||
buffer<expr> f_types, new_f_types; // used only if m_proofs_enabled
|
||||
buffer<bool> heq_proofs; // used only if m_has_heq && m_proofs_enabled
|
||||
buffer<bool> heq_proofs; // used only if m_use_heq && m_proofs_enabled
|
||||
bool changed = false;
|
||||
expr f = arg(e, 0);
|
||||
expr f_type = infer_type(f);
|
||||
|
@ -643,7 +649,7 @@ class simplifier_cell::imp {
|
|||
f_types.push_back(f_type);
|
||||
new_f_type = res_f.is_heq_proof() ? infer_type(new_f) : f_type;
|
||||
new_f_types.push_back(new_f_type);
|
||||
if (m_has_heq)
|
||||
if (m_use_heq)
|
||||
heq_proofs.push_back(res_f.is_heq_proof());
|
||||
}
|
||||
unsigned num = num_args(e);
|
||||
|
@ -652,7 +658,7 @@ class simplifier_cell::imp {
|
|||
bool f_arrow = is_arrow(f_type);
|
||||
expr const & a = arg(e, i);
|
||||
result res_a(a);
|
||||
if (m_has_heq || f_arrow) {
|
||||
if (m_use_heq || f_arrow) {
|
||||
res_a = simplify(a);
|
||||
if (res_a.m_expr != a)
|
||||
changed = true;
|
||||
|
@ -661,7 +667,7 @@ class simplifier_cell::imp {
|
|||
new_args.push_back(new_a);
|
||||
if (m_proofs_enabled) {
|
||||
proofs.push_back(res_a.m_proof);
|
||||
if (m_has_heq)
|
||||
if (m_use_heq)
|
||||
heq_proofs.push_back(res_a.is_heq_proof());
|
||||
bool changed_f_type = !is_eqp(f_type, new_f_type);
|
||||
if (f_arrow) {
|
||||
|
@ -696,8 +702,8 @@ class simplifier_cell::imp {
|
|||
bool heq_proof = false;
|
||||
if (i == 0) {
|
||||
pr = *(proofs[0]);
|
||||
heq_proof = m_has_heq && heq_proofs[0];
|
||||
} else if (m_has_heq && (heq_proofs[i] || !is_arrow(f_types[i-1]))) {
|
||||
heq_proof = m_use_heq && heq_proofs[0];
|
||||
} else if (m_use_heq && (heq_proofs[i] || !is_arrow(f_types[i-1]))) {
|
||||
expr f = mk_app_prefix(i, new_args);
|
||||
expr pr_i = *proofs[i];
|
||||
auto new_pr = mk_hcongr_th(f_types[i-1], f_types[i-1], f, f, mk_hrefl_th(f_types[i-1], f),
|
||||
|
@ -717,7 +723,7 @@ class simplifier_cell::imp {
|
|||
expr new_f = mk_app_prefix(i, new_args);
|
||||
if (proofs[i]) {
|
||||
expr pr_i = *proofs[i];
|
||||
if (m_has_heq && heq_proofs[i]) {
|
||||
if (m_use_heq && heq_proofs[i]) {
|
||||
if (!heq_proof)
|
||||
pr = mk_to_heq_th(f_types[i], f, new_f, pr);
|
||||
auto new_pr = mk_hcongr_th(f_types[i-1], new_f_types[i-1], f, new_f, pr,
|
||||
|
@ -1125,7 +1131,7 @@ class simplifier_cell::imp {
|
|||
if (!m_proofs_enabled || !res_bi.m_proof)
|
||||
return rewrite_lambda(e, result(new_e));
|
||||
if (res_bi.is_heq_proof()) {
|
||||
lean_assert(m_has_heq);
|
||||
lean_assert(m_use_heq);
|
||||
// Using
|
||||
// theorem hsfunext {A : TypeM} {B B' : A → TypeU} {f : ∀ x, B x} {f' : ∀ x, B' x} :
|
||||
// (∀ x, f x == f' x) → f == f'
|
||||
|
@ -1200,7 +1206,7 @@ class simplifier_cell::imp {
|
|||
|
||||
result simplify_lambda(expr const & e) {
|
||||
lean_assert(is_lambda(e));
|
||||
if (m_has_heq) {
|
||||
if (m_use_heq) {
|
||||
return simplify_lambda_with_heq(e);
|
||||
} else {
|
||||
return simplify_lambda_body(e);
|
||||
|
@ -1427,7 +1433,7 @@ class simplifier_cell::imp {
|
|||
else
|
||||
return simplify_arrow(e);
|
||||
} else if (is_proposition(e)) {
|
||||
if (m_has_heq)
|
||||
if (m_use_heq)
|
||||
return simplify_forall_with_heq(e);
|
||||
else
|
||||
return simplify_forall_body(e);
|
||||
|
@ -1505,6 +1511,7 @@ class simplifier_cell::imp {
|
|||
m_conditional = get_simplifier_conditional(o);
|
||||
m_memoize = get_simplifier_memoize(o);
|
||||
m_max_steps = get_simplifier_max_steps(o);
|
||||
m_use_heq = get_simplifier_heq(o);
|
||||
m_preserve_binder_names = get_simplifier_preserve_binder_names(o);
|
||||
}
|
||||
|
||||
|
@ -1512,7 +1519,6 @@ public:
|
|||
imp(ro_environment const & env, options const & o, unsigned num_rs, rewrite_rule_set const * rs,
|
||||
std::shared_ptr<simplifier_monitor> const & monitor):
|
||||
m_env(env), m_options(o), m_tc(env), m_monitor(monitor) {
|
||||
m_has_heq = m_env->imported("heq");
|
||||
set_options(o);
|
||||
if (m_contextual) {
|
||||
// We need an extra rule set if we are performing contextual rewriting
|
||||
|
|
|
@ -9,7 +9,6 @@ Author: Leonardo de Moura
|
|||
#include "util/sexpr/option_declarations.h"
|
||||
#include "kernel/type_checker.h"
|
||||
#include "kernel/kernel.h"
|
||||
#include "library/heq_decls.h"
|
||||
#include "library/simplifier/simplifier.h"
|
||||
#include "library/tactic/tactic.h"
|
||||
|
||||
|
|
|
@ -1,4 +1,3 @@
|
|||
import heq
|
||||
import tactic
|
||||
add_rewrite Nat::add_assoc Nat::add_zeror
|
||||
|
||||
|
@ -14,9 +13,10 @@ axiom concat_empty {n : Nat} (v : vec n) :
|
|||
add_rewrite concat_assoc concat_empty
|
||||
|
||||
(*
|
||||
local opts = options({"simplifier", "heq"}, true)
|
||||
local t = parse_lean('∀ (n : Nat) (v : vec (n + 0)) (w : vec n), v = w ; empty')
|
||||
print(t)
|
||||
local t2, pr = simplify(t)
|
||||
local t2, pr = simplify(t, "default", opts)
|
||||
print("====>")
|
||||
print(t2)
|
||||
get_environment():type_check(pr)
|
||||
|
@ -25,9 +25,10 @@ get_environment():type_check(pr)
|
|||
print ""
|
||||
|
||||
(*
|
||||
local opts = options({"simplifier", "heq"}, true)
|
||||
local t = parse_lean('λ n : Nat, ∃ (v : vec (n + 0)) (w : vec n), v ≠ w ; empty')
|
||||
print(t)
|
||||
local t2, pr = simplify(t)
|
||||
local t2, pr = simplify(t, "default", opts)
|
||||
print("====>")
|
||||
print(t2)
|
||||
get_environment():type_check(pr)
|
||||
|
|
|
@ -1,6 +1,5 @@
|
|||
Set: pp::colors
|
||||
Set: pp::unicode
|
||||
Imported 'heq'
|
||||
Imported 'tactic'
|
||||
Assumed: vec
|
||||
Assumed: concat
|
||||
|
|
|
@ -1,4 +1,3 @@
|
|||
import heq
|
||||
variable vec : Nat → Type
|
||||
variable concat {n m : Nat} (v : vec n) (w : vec m) : vec (n + m)
|
||||
infixl 65 ; : concat
|
||||
|
@ -18,10 +17,11 @@ variable v : vec n
|
|||
variable w : vec n
|
||||
|
||||
(*
|
||||
local opts = options({"simplifier", "heq"}, true)
|
||||
local t = parse_lean([[ (v ; w) ; empty ; (v ; empty) ]])
|
||||
print(t)
|
||||
print("===>")
|
||||
local t2, pr = simplify(t, "simple")
|
||||
local t2, pr = simplify(t, "simple", opts)
|
||||
print(t2)
|
||||
print(pr)
|
||||
get_environment():type_check(pr)
|
||||
|
|
|
@ -1,6 +1,5 @@
|
|||
Set: pp::colors
|
||||
Set: pp::unicode
|
||||
Imported 'heq'
|
||||
Assumed: vec
|
||||
Assumed: concat
|
||||
Assumed: concat_assoc
|
||||
|
|
|
@ -1,4 +1,3 @@
|
|||
import heq
|
||||
variable vec : Nat → Type
|
||||
variable concat {n m : Nat} (v : vec n) (w : vec m) : vec (n + m)
|
||||
infixl 65 ; : concat
|
||||
|
@ -19,10 +18,11 @@ variable w : vec n
|
|||
variable f {A : TypeM} : A → A
|
||||
|
||||
(*
|
||||
local opts = options({"simplifier", "heq"}, true)
|
||||
local t = parse_lean([[ f ((v ; w) ; empty ; (v ; empty)) ]])
|
||||
print(t)
|
||||
print("===>")
|
||||
local t2, pr = simplify(t, "simple")
|
||||
local t2, pr = simplify(t, "simple", opts)
|
||||
print(t2)
|
||||
print(pr)
|
||||
get_environment():type_check(pr)
|
||||
|
|
|
@ -1,6 +1,5 @@
|
|||
Set: pp::colors
|
||||
Set: pp::unicode
|
||||
Imported 'heq'
|
||||
Assumed: vec
|
||||
Assumed: concat
|
||||
Assumed: concat_assoc
|
||||
|
|
|
@ -1,4 +1,3 @@
|
|||
import heq
|
||||
variable vec : Nat → Type
|
||||
variable concat {n m : Nat} (v : vec n) (w : vec m) : vec (n + m)
|
||||
infixl 65 ; : concat
|
||||
|
@ -22,10 +21,11 @@ variable w : vec n
|
|||
variable f {A : TypeM} : A → A
|
||||
|
||||
(*
|
||||
local opts = options({"simplifier", "heq"}, true)
|
||||
local t = parse_lean([[ f ((v ; w) ; empty ; (v ; empty)) ]])
|
||||
print(t)
|
||||
print("===>")
|
||||
local t2, pr = simplify(t, "simple")
|
||||
local t2, pr = simplify(t, "simple", opts)
|
||||
print(t2)
|
||||
print(pr)
|
||||
get_environment():type_check(pr)
|
||||
|
@ -36,10 +36,11 @@ disable_rewrite Nat::add_comm : simple
|
|||
print "After disabling Nat::add_comm"
|
||||
|
||||
(*
|
||||
local opts = options({"simplifier", "heq"}, true)
|
||||
local t = parse_lean([[ f ((v ; w) ; empty ; (v ; empty)) ]])
|
||||
print(t)
|
||||
print("===>")
|
||||
local t2, pr = simplify(t, "simple")
|
||||
local t2, pr = simplify(t, "simple", opts)
|
||||
print(t2)
|
||||
print(pr)
|
||||
get_environment():type_check(pr)
|
||||
|
|
|
@ -1,6 +1,5 @@
|
|||
Set: pp::colors
|
||||
Set: pp::unicode
|
||||
Imported 'heq'
|
||||
Assumed: vec
|
||||
Assumed: concat
|
||||
Assumed: concat_assoc
|
||||
|
|
|
@ -1,4 +1,3 @@
|
|||
import heq
|
||||
variable vec : Nat → Type
|
||||
variable concat {n m : Nat} (v : vec n) (w : vec m) : vec (n + m)
|
||||
infixl 65 ; : concat
|
||||
|
@ -22,10 +21,11 @@ axiom fax {n m : Nat} (v : vec n) (w : vec m) : f (v; (w; v)) = v; (w; v)
|
|||
add_rewrite fax : simple
|
||||
|
||||
(*
|
||||
local opts = options({"simplifier", "heq"}, true)
|
||||
local t = parse_lean([[ p (f ((v ; w) ; empty ; (v ; empty))) ∧ v = cast (congr2 vec (Nat::add_zeror n)) (v ; empty) ]])
|
||||
print(t)
|
||||
print("===>")
|
||||
local t2, pr = simplify(t, "simple")
|
||||
local t2, pr = simplify(t, "simple", opts)
|
||||
print(t2)
|
||||
print("checking proof")
|
||||
print (get_environment():type_check(pr))
|
||||
|
|
|
@ -1,6 +1,5 @@
|
|||
Set: pp::colors
|
||||
Set: pp::unicode
|
||||
Imported 'heq'
|
||||
Assumed: vec
|
||||
Assumed: concat
|
||||
Assumed: concat_assoc
|
||||
|
|
|
@ -1,4 +1,3 @@
|
|||
import heq
|
||||
variable vec : Nat → Type
|
||||
variable concat {n m : Nat} (v : vec n) (w : vec m) : vec (n + m)
|
||||
infixl 65 ; : concat
|
||||
|
@ -15,8 +14,9 @@ add_rewrite Nat::add_assoc Nat::add_zeror eq_id : simple
|
|||
add_rewrite concat_assoc concat_empty Nat::add_assoc Nat::add_zeror : simple
|
||||
|
||||
(*
|
||||
local opts = options({"simplifier", "heq"}, true)
|
||||
local t = parse_lean('λ n : Nat, λ v : vec n, v ; empty')
|
||||
local t2, pr = simplify(t, "simple")
|
||||
local t2, pr = simplify(t, "simple", opts)
|
||||
print(t2)
|
||||
-- print(pr)
|
||||
get_environment():type_check(pr)
|
||||
|
@ -25,8 +25,9 @@ get_environment():type_check(pr)
|
|||
variable f {A : Type} : A → A
|
||||
|
||||
(*
|
||||
local opts = options({"simplifier", "heq"}, true)
|
||||
local t = parse_lean('λ n : Nat, λ v : vec (n + 0), (f v) ; empty')
|
||||
local t2, pr = simplify(t, "simple")
|
||||
local t2, pr = simplify(t, "simple", opts)
|
||||
print(t2)
|
||||
-- print(pr)
|
||||
get_environment():type_check(pr)
|
||||
|
@ -37,10 +38,11 @@ print ""
|
|||
variable lheq {A B : TypeM} : A → B → Bool
|
||||
infixl 50 === : lheq
|
||||
(*
|
||||
local opts = options({"simplifier", "heq"}, true)
|
||||
local t = parse_lean('λ val : Nat, (λ n : Nat, λ v : vec (n + 0), (f v) ; empty) val === (λ n : Nat, λ v : vec (n + 0), v) val')
|
||||
print(t)
|
||||
print("=====>")
|
||||
local t2, pr = simplify(t, "simple")
|
||||
local t2, pr = simplify(t, "simple", opts)
|
||||
print(t2)
|
||||
-- print(pr)
|
||||
get_environment():type_check(pr)
|
||||
|
|
|
@ -1,6 +1,5 @@
|
|||
Set: pp::colors
|
||||
Set: pp::unicode
|
||||
Imported 'heq'
|
||||
Assumed: vec
|
||||
Assumed: concat
|
||||
Assumed: concat_assoc
|
||||
|
|
|
@ -1,4 +1,3 @@
|
|||
import heq
|
||||
variable vec : Nat → Type
|
||||
variable concat {n m : Nat} (v : vec n) (w : vec m) : vec (n + m)
|
||||
infixl 65 ; : concat
|
||||
|
@ -15,9 +14,10 @@ add_rewrite Nat::add_assoc Nat::add_zeror eq_id : simple
|
|||
add_rewrite concat_assoc concat_empty Nat::add_assoc Nat::add_zeror : simple
|
||||
|
||||
(*
|
||||
local opts = options({"simplifier", "heq"}, true)
|
||||
local t = parse_lean('∀ (n : Nat) (v : vec (n + 0)) (w : vec n), v = w ; empty')
|
||||
print(t)
|
||||
local t2, pr = simplify(t, "simple")
|
||||
local t2, pr = simplify(t, "simple", opts)
|
||||
print("====>")
|
||||
print(t2)
|
||||
get_environment():type_check(pr)
|
||||
|
@ -26,9 +26,10 @@ get_environment():type_check(pr)
|
|||
print ""
|
||||
|
||||
(*
|
||||
local opts = options({"simplifier", "heq"}, true)
|
||||
local t = parse_lean('λ n : Nat, ∃ (v : vec (n + 0)) (w : vec n), v ≠ w ; empty')
|
||||
print(t)
|
||||
local t2, pr = simplify(t, "simple")
|
||||
local t2, pr = simplify(t, "simple", opts)
|
||||
print("====>")
|
||||
print(t2)
|
||||
get_environment():type_check(pr)
|
||||
|
|
|
@ -1,6 +1,5 @@
|
|||
Set: pp::colors
|
||||
Set: pp::unicode
|
||||
Imported 'heq'
|
||||
Assumed: vec
|
||||
Assumed: concat
|
||||
Assumed: concat_assoc
|
||||
|
|
|
@ -1,4 +1,3 @@
|
|||
import heq
|
||||
variable vec : Nat → Type
|
||||
variable concat {n m : Nat} (v : vec n) (w : vec m) : vec (n + m)
|
||||
infixl 65 ; : concat
|
||||
|
@ -17,6 +16,7 @@ add_rewrite concat_assoc concat_empty Nat::add_assoc Nat::add_zeror : simple
|
|||
variable f {A : Type} : A → A
|
||||
|
||||
(*
|
||||
local opts = options({"simplifier", "heq"}, true)
|
||||
local m = simplifier_monitor(nil, nil, nil,
|
||||
function (s, e, i, k)
|
||||
print("App simplification failure, argument #" .. i)
|
||||
|
@ -30,7 +30,7 @@ local m = simplifier_monitor(nil, nil, nil,
|
|||
print("Abst failure: " .. tostring(e))
|
||||
end
|
||||
)
|
||||
local s = simplifier("simple", options(), m)
|
||||
local s = simplifier("simple", opts, m)
|
||||
local t = parse_lean('λ val : Nat, (λ n : Nat, λ v : vec (n + 0), (f v) ; empty) val == (λ n : Nat, λ v : vec (n + 0), v) val')
|
||||
print(t)
|
||||
print("=====>")
|
||||
|
|
|
@ -1,6 +1,5 @@
|
|||
Set: pp::colors
|
||||
Set: pp::unicode
|
||||
Imported 'heq'
|
||||
Assumed: vec
|
||||
Assumed: concat
|
||||
Assumed: concat_assoc
|
||||
|
|
|
@ -1,5 +1,5 @@
|
|||
import tactic
|
||||
import heq
|
||||
set_option simplifier::heq true
|
||||
set_option pp::implicit true -- to be able to parse output produced by Lean
|
||||
variable vec : Nat → Type
|
||||
variables n m : Nat
|
||||
|
|
|
@ -1,7 +1,7 @@
|
|||
Set: pp::colors
|
||||
Set: pp::unicode
|
||||
Imported 'tactic'
|
||||
Imported 'heq'
|
||||
Set: simplifier::heq
|
||||
Set: lean::pp::implicit
|
||||
Assumed: vec
|
||||
Assumed: n
|
||||
|
|
Loading…
Reference in a new issue