feat(library): add 'noncomputable' keyword for the standard library

This commit is contained in:
Leonardo de Moura 2015-07-28 21:56:35 -07:00
parent a009db2902
commit 308af87b69
33 changed files with 166 additions and 72 deletions

View file

@ -329,6 +329,8 @@ section discrete_linear_ordered_field
(assume H' : ¬ y < x,
decidable.inl (le.antisymm (le_of_not_gt H') (le_of_not_gt H))))
reveal dec_eq_of_dec_lt
definition discrete_linear_ordered_field.to_discrete_field [trans-instance] [reducible] [coercion]
: discrete_field A :=
⦃ discrete_field, s, has_decidable_eq := dec_eq_of_dec_lt⦄

View file

@ -13,6 +13,8 @@ namespace empty
subsingleton.intro (λ a b, !empty.elim a)
end empty
reveal empty.elim
protected definition empty.has_decidable_eq [instance] : decidable_eq empty :=
take (a b : empty), decidable.inl (!empty.elim a)

View file

@ -14,7 +14,7 @@ open relation nonempty subtype
/- abstract quotient -/
definition prelim_map {A : Type} (R : A → A → Prop) (a : A) :=
noncomputable definition prelim_map {A : Type} (R : A → A → Prop) (a : A) :=
-- TODO: it is interesting how the elaborator fails here
-- epsilon (fun b, R a b)
@epsilon _ (nonempty.intro a) (fun b, R a b)
@ -42,12 +42,12 @@ assert H5 : nonempty A, from
show epsilon (λc, R a c) = epsilon (λc, R b c), from
congr_arg _ H4
definition quotient {A : Type} (R : A → A → Prop) : Type := image (prelim_map R)
noncomputable definition quotient {A : Type} (R : A → A → Prop) : Type := image (prelim_map R)
definition quotient_abs {A : Type} (R : A → A → Prop) : A → quotient R :=
noncomputable definition quotient_abs {A : Type} (R : A → A → Prop) : A → quotient R :=
fun_image (prelim_map R)
definition quotient_elt_of {A : Type} (R : A → A → Prop) : quotient R → A := elt_of
noncomputable definition quotient_elt_of {A : Type} (R : A → A → Prop) : quotient R → A := elt_of
theorem quotient_is_quotient {A : Type} (R : A → A → Prop) (H : is_equivalence R)
: is_quotient R (quotient_abs R) (quotient_elt_of R) :=

View file

@ -216,9 +216,9 @@ theorem rewrite_helper9 (a b c : ) : b - c = (b - a) - (c - a) :=
theorem rewrite_helper10 (a b c d : ) : c - d = (c - a) + (a - b) + (b - d) :=
by rewrite[*add_sub,*sub_add_cancel]
definition rep (x : ) : s.reg_seq := some (quot.exists_rep x)
noncomputable definition rep (x : ) : s.reg_seq := some (quot.exists_rep x)
definition re_abs (x : ) : :=
noncomputable definition re_abs (x : ) : :=
quot.lift_on x (λ a, quot.mk (s.r_abs a)) (take a b Hab, quot.sound (s.r_abs_well_defined Hab))
theorem r_abs_nonneg {x : } : zero ≤ x → re_abs x = x :=
@ -251,7 +251,7 @@ theorem rat_approx' (x : ) : ∀ n : +, ∃ q : , re_abs (x - const q)
theorem rat_approx (x : ) : ∀ n : +, ∃ q : , abs (x - const q) ≤ const n⁻¹ :=
by rewrite -re_abs_is_abs; apply rat_approx'
definition approx (x : ) (n : +) := some (rat_approx x n)
noncomputable definition approx (x : ) (n : +) := some (rat_approx x n)
theorem approx_spec (x : ) (n : +) : abs (x - (const (approx x n))) ≤ const n⁻¹ :=
some_spec (rat_approx x n)
@ -261,10 +261,10 @@ theorem approx_spec' (x : ) (n : +) : abs ((const (approx x n)) - x) ≤ c
notation `r_seq` := + →
definition converges_to (X : r_seq) (a : ) (N : + → +) :=
noncomputable definition converges_to (X : r_seq) (a : ) (N : + → +) :=
∀ k : +, ∀ n : +, n ≥ N k → abs (X n - a) ≤ const k⁻¹
definition cauchy (X : r_seq) (M : + → +) :=
noncomputable definition cauchy (X : r_seq) (M : + → +) :=
∀ k : +, ∀ m n : +, m ≥ M k → n ≥ M k → abs (X m - X n) ≤ const k⁻¹
--set_option pp.implicit true
--set_option pp.coercions true
@ -312,7 +312,7 @@ theorem Nb_spec_right (M : + → +) (k : +) : M (2 * k) ≤ Nb M k := !
theorem Nb_spec_left (M : + → +) (k : +) : 3 * k ≤ Nb M k := !max_left
definition lim_seq {X : r_seq} {M : + → +} (Hc : cauchy X M) : + → :=
noncomputable definition lim_seq {X : r_seq} {M : + → +} (Hc : cauchy X M) : + → :=
λ k, approx (X (Nb M k)) (2 * k)
theorem lim_seq_reg_helper {X : r_seq} {M : + → +} (Hc : cauchy X M) {m n : +}
@ -368,14 +368,14 @@ theorem lim_seq_spec {X : r_seq} {M : + → +} (Hc : cauchy X M) (k : +
apply lim_seq_reg
end
definition r_lim_seq {X : r_seq} {M : + → +} (Hc : cauchy X M) : s.reg_seq :=
noncomputable definition r_lim_seq {X : r_seq} {M : + → +} (Hc : cauchy X M) : s.reg_seq :=
s.reg_seq.mk (lim_seq Hc) (lim_seq_reg Hc)
theorem r_lim_seq_spec {X : r_seq} {M : + → +} (Hc : cauchy X M) (k : +) :
s.r_le (s.r_abs (( s.radd (r_lim_seq Hc) (s.rneg (s.r_const ((s.reg_seq.sq (r_lim_seq Hc)) k)))))) (s.r_const (k)⁻¹) :=
lim_seq_spec Hc k
definition lim {X : r_seq} {M : + → +} (Hc : cauchy X M) : :=
noncomputable definition lim {X : r_seq} {M : + → +} (Hc : cauchy X M) : :=
quot.mk (r_lim_seq Hc)
theorem re_lim_spec {x : r_seq} {M : + → +} (Hc : cauchy x M) (k : +) :

View file

@ -107,9 +107,9 @@ theorem sep_zero_of_pos {s : seq} (Hs : regular s) (Hpos : pos s) : sep s zero :
------------------------
-- This section could be cleaned up.
definition pb {s : seq} (Hs : regular s) (Hpos : pos s) :=
noncomputable definition pb {s : seq} (Hs : regular s) (Hpos : pos s) :=
some (abs_pos_of_nonzero Hs (sep_zero_of_pos Hs Hpos))
definition ps {s : seq} (Hs : regular s) (Hsep : sep s zero) :=
noncomputable definition ps {s : seq} (Hs : regular s) (Hsep : sep s zero) :=
some (abs_pos_of_nonzero Hs Hsep)
@ -121,7 +121,7 @@ theorem ps_spec {s : seq} (Hs : regular s) (Hsep : sep s zero) :
∀ m : +, m ≥ (ps Hs Hsep) → abs (s m) ≥ (ps Hs Hsep)⁻¹ :=
some_spec (abs_pos_of_nonzero Hs Hsep)
definition s_inv {s : seq} (Hs : regular s) (n : +) : :=
noncomputable definition s_inv {s : seq} (Hs : regular s) (n : +) : :=
if H : sep s zero then
(if n < (ps Hs H) then 1 / (s ((ps Hs H) * (ps Hs H) * (ps Hs H)))
else 1 / (s ((ps Hs H) * (ps Hs H) * n)))
@ -570,7 +570,7 @@ theorem s_le_of_equiv_le_right {s t u : seq} (Hs : regular s) (Ht : regular t) (
-----------------------------
definition r_inv (s : reg_seq) : reg_seq := reg_seq.mk (s_inv (reg_seq.is_reg s))
noncomputable definition r_inv (s : reg_seq) : reg_seq := reg_seq.mk (s_inv (reg_seq.is_reg s))
(if H : sep (reg_seq.sq s) zero then reg_inv_reg (reg_seq.is_reg s) H else
have Hz : s_inv (reg_seq.is_reg s) = zero, from funext (λ n, dif_neg H), Hz⁻¹ ▸ zero_is_reg)
@ -605,7 +605,7 @@ end s
namespace real
open [classes] s
definition inv (x : ) : := quot.lift_on x (λ a, quot.mk (s.r_inv a))
noncomputable definition inv (x : ) : := quot.lift_on x (λ a, quot.mk (s.r_inv a))
(λ a b H, quot.sound (s.r_inv_well_defined H))
postfix [priority real.prio] `⁻¹` := inv
@ -651,7 +651,7 @@ theorem dec_lt : decidable_rel lt :=
section migrate_algebra
open [classes] algebra
protected definition discrete_linear_ordered_field [reducible] :
protected noncomputable definition discrete_linear_ordered_field [reducible] :
algebra.discrete_linear_ordered_field :=
⦃ algebra.discrete_linear_ordered_field, real.comm_ring, real.ordered_ring,
le_total := le_total,
@ -667,11 +667,11 @@ section migrate_algebra
local attribute real.comm_ring [instance]
local attribute real.ordered_ring [instance]
definition abs (n : ) : := algebra.abs n
definition sign (n : ) : := algebra.sign n
definition max (a b : ) : := algebra.max a b
definition min (a b : ) : := algebra.min a b
definition divide (a b : ): := algebra.divide a b
noncomputable definition abs (n : ) : := algebra.abs n
noncomputable definition sign (n : ) : := algebra.sign n
noncomputable definition max (a b : ) : := algebra.max a b
noncomputable definition min (a b : ) : := algebra.min a b
noncomputable definition divide (a b : ): := algebra.divide a b
migrate from algebra with real
replacing has_le.ge → ge, has_lt.gt → gt, sub → sub, abs → abs, sign → sign, dvd → dvd,

View file

@ -13,7 +13,7 @@ namespace set
variables {X Y : Type}
definition inv_fun (f : X → Y) (a : set X) (dflt : X) (y : Y) : X :=
noncomputable definition inv_fun (f : X → Y) (a : set X) (dflt : X) (y : Y) : X :=
if H : ∃₀ x ∈ a, f x = y then some H else dflt
theorem inv_fun_pos {f : X → Y} {a : set X} {dflt : X} {y : Y}
@ -65,7 +65,7 @@ namespace map
variables {X Y : Type} {a : set X} {b : set Y}
protected definition inverse (f : map a b) {dflt : X} (dflta : dflt ∈ a) :=
protected noncomputable definition inverse (f : map a b) {dflt : X} (dflta : dflt ∈ a) :=
map.mk (inv_fun f a dflt) (@maps_to_inv_fun _ _ _ _ b _ dflta)
theorem left_inverse_inverse {f : map a b} {dflt : X} (dflta : dflt ∈ a) (H : map.injective f) :

View file

@ -18,7 +18,7 @@ definition trivial := true.intro
definition not (a : Prop) := a → false
prefix `¬` := not
theorem absurd {a : Prop} {b : Type} (H1 : a) (H2 : ¬a) : b :=
definition absurd {a : Prop} {b : Type} (H1 : a) (H2 : ¬a) : b :=
false.rec b (H2 H1)
/- not -/

View file

@ -196,7 +196,7 @@ namespace nat
(λh, inl (le_succ_of_le h))
(λh, elim (eq_or_lt_of_le h) (λe, inl (eq.subst e !le.refl)) inr)) b
theorem lt_ge_by_cases {a b : } {P : Type} (H1 : a < b → P) (H2 : a ≥ b → P) : P :=
definition lt_ge_by_cases {a b : } {P : Type} (H1 : a < b → P) (H2 : a ≥ b → P) : P :=
by_cases H1 (λh, H2 (elim !lt_or_ge (λa, absurd a h) (λa, a)))
definition lt.by_cases {a b : } {P : Type} (H1 : a < b → P) (H2 : a = b → P) (H3 : b < a → P) : P :=

View file

@ -18,8 +18,8 @@ parameter p : Prop
private definition U (x : Prop) : Prop := x = true p
private definition V (x : Prop) : Prop := x = false p
private definition u := epsilon U
private definition v := epsilon V
private noncomputable definition u := epsilon U
private noncomputable definition v := epsilon V
private lemma u_def : U u :=
epsilon_spec (exists.intro true (or.inl rfl))

View file

@ -12,7 +12,7 @@ with (∃ a : A, f a = b).
import logic.axioms.classical
open function
definition mk_left_inv {A B : Type} [h : nonempty A] (f : A → B) : B → A :=
noncomputable definition mk_left_inv {A B : Type} [h : nonempty A] (f : A → B) : B → A :=
λ b : B, if ex : (∃ a : A, f a = b) then some ex else inhabited.value (inhabited_of_nonempty h)
theorem has_left_inverse_of_injective {A B : Type} {f : A → B} : nonempty A → injective f → has_left_inverse f :=

View file

@ -28,7 +28,7 @@ inhabited_of_nonempty (obtain w Hw, from H, nonempty.intro w)
/- the Hilbert epsilon function -/
definition epsilon {A : Type} [H : nonempty A] (P : A → Prop) : A :=
noncomputable definition epsilon {A : Type} [H : nonempty A] (P : A → Prop) : A :=
let u : {x | (∃y, P y) → P x} :=
strong_indefinite_description P H in
elt_of u
@ -47,7 +47,7 @@ epsilon_spec_aux (nonempty_of_exists Hex) P Hex
theorem epsilon_singleton {A : Type} (a : A) : @epsilon A (nonempty.intro a) (λx, x = a) = a :=
epsilon_spec (exists.intro a (eq.refl a))
definition some {A : Type} {P : A → Prop} (H : ∃x, P x) : A :=
noncomputable definition some {A : Type} {P : A → Prop} (H : ∃x, P x) : A :=
@epsilon A (nonempty_of_exists H) P
theorem some_spec {A : Type} {P : A → Prop} (H : ∃x, P x) : P (some H) :=

View file

@ -32,7 +32,7 @@ axiom recA : Π {C : A → Type}, (Π (p : Phi A), C (introA p)) → (Π (a :
axiom recA_comp : Π {C : A → Type} (H : Π (p : Phi A), C (introA p)) (p : Phi A), recA H (introA p) = H p
-- The recursor could be used to define matchA
definition matchA (a : A) : Phi A :=
noncomputable definition matchA (a : A) : Phi A :=
recA (λ p, p) a
-- and the computation rule would allows us to define
@ -57,7 +57,7 @@ lemma i_injective {T : Type} {a b : T} : i a = i b → a = b :=
eq.symm e₃
-- Hence, by composition, we get an injection f from (A → Prop) to A
definition f : (A → Prop) → A :=
noncomputable definition f : (A → Prop) → A :=
λ p, introA (i p)
lemma f_injective : ∀ {p p' : A → Prop}, f p = f p' → p = p':=
@ -68,11 +68,11 @@ lemma f_injective : ∀ {p p' : A → Prop}, f p = f p' → p = p':=
We are now back to the usual Cantor-Russel paradox.
We can define
-/
definition P0 (a : A) : Prop :=
noncomputable definition P0 (a : A) : Prop :=
∃ (P : A → Prop), f P = a ∧ ¬ P a
-- i.e., P0 a := codes a set P such that x∉P
definition x0 : A := f P0
noncomputable definition x0 : A := f P0
lemma fP0_eq : f P0 = x0 :=
rfl

View file

@ -23,8 +23,8 @@ axiom introD : (D → D) → D
axiom recD : Π {C : D → Type}, (Π (f : D → D) (r : Π d, C (f d)), C (introD f)) → (Π (d : D), C d)
-- We would also get a computational rule for the eliminator, but we don't need it for deriving the inconsistency.
definition id : D → D := λd, d
definition v : D := introD id
noncomputable definition id : D → D := λd, d
noncomputable definition v : D := introD id
theorem inconsistent : false :=
recD (λ f ih, ih v) v

View file

@ -35,6 +35,8 @@ is_congruence2.mk @congr_imp
theorem is_congruence_iff : is_congruence2 iff iff iff iff :=
is_congruence2.mk @congr_iff
reveal is_congruence_not is_congruence_and is_congruence_or is_congruence_imp is_congruence_iff
definition is_congruence_not_compose [instance] := is_congruence.compose is_congruence_not
definition is_congruence_and_compose [instance] := is_congruence.compose21 is_congruence_and
definition is_congruence_or_compose [instance] := is_congruence.compose21 is_congruence_or

View file

@ -8,7 +8,7 @@
(require 'rx)
(defconst lean-keywords
'("import" "prelude" "tactic_hint" "protected" "private" "definition" "renaming"
'("import" "prelude" "tactic_hint" "protected" "private" "noncomputable" "definition" "renaming"
"hiding" "exposing" "parameter" "parameters" "begin" "begin+" "proof" "qed" "conjecture" "constant" "constants"
"hypothesis" "lemma" "corollary" "variable" "variables" "premise" "premises"
"print" "theorem" "example" "abbreviation" "abstract"

View file

@ -16,6 +16,7 @@ Author: Leonardo de Moura
#include "library/aliases.h"
#include "library/private.h"
#include "library/protected.h"
#include "library/noncomputable.h"
#include "library/placeholder.h"
#include "library/locals.h"
#include "library/explicit.h"
@ -630,6 +631,7 @@ class definition_cmd_fn {
def_cmd_kind m_kind;
bool m_is_private;
bool m_is_protected;
bool m_is_noncomputable;
pos_info m_pos;
name m_name;
@ -864,6 +866,16 @@ class definition_cmd_fn {
void register_decl(name const & n, name const & real_n, expr const & type) {
if (m_kind != Example) {
if (!m_is_noncomputable && is_marked_noncomputable(m_env, real_n)) {
auto reason = get_noncomputable_reason(m_env, real_n);
lean_assert(reason);
if (m_p.in_theorem_queue(*reason)) {
throw exception(sstream() << "definition '" << n << "' was marked as noncomputable because '" << *reason
<< "' is still in theorem queue (solution: use command 'reveal " << *reason << "'");
} else {
throw exception(sstream() << "definition '" << n << "' is noncomputable, it depends on '" << *reason << "'");
}
}
// TODO(Leo): register aux_decls
if (!m_is_private)
m_p.add_decl_index(real_n, m_pos, m_p.get_cmd_token(), type);
@ -995,7 +1007,7 @@ class definition_cmd_fn {
m_ls = append(m_ls, new_ls);
m_type = postprocess(m_env, m_type);
expr type_as_is = m_p.save_pos(mk_as_is(m_type), type_pos);
if (!m_p.collecting_info() && m_kind == Theorem && m_p.num_threads() > 1) {
if (!m_p.collecting_info() && !m_is_noncomputable && m_kind == Theorem && m_p.num_threads() > 1) {
// Add as axiom, and create a task to prove the theorem.
// Remark: we don't postpone the "proof" of Examples.
m_p.add_delayed_theorem(m_env, m_real_name, m_ls, type_as_is, m_value);
@ -1053,9 +1065,9 @@ class definition_cmd_fn {
}
public:
definition_cmd_fn(parser & p, def_cmd_kind kind, bool is_private, bool is_protected):
definition_cmd_fn(parser & p, def_cmd_kind kind, bool is_private, bool is_protected, bool is_noncomputable):
m_p(p), m_env(m_p.env()), m_kind(kind),
m_is_private(is_private), m_is_protected(is_protected),
m_is_private(is_private), m_is_protected(is_protected), m_is_noncomputable(is_noncomputable),
m_pos(p.pos()), m_attributes(kind == Abbreviation || kind == LocalAbbreviation) {
lean_assert(!(m_is_private && m_is_protected));
}
@ -1082,26 +1094,31 @@ public:
}
};
static environment definition_cmd_core(parser & p, def_cmd_kind kind, bool is_private, bool is_protected) {
return definition_cmd_fn(p, kind, is_private, is_protected)();
static environment definition_cmd_core(parser & p, def_cmd_kind kind, bool is_private, bool is_protected, bool is_noncomputable) {
return definition_cmd_fn(p, kind, is_private, is_protected, is_noncomputable)();
}
static environment definition_cmd(parser & p) {
return definition_cmd_core(p, Definition, false, false);
return definition_cmd_core(p, Definition, false, false, false);
}
static environment abbreviation_cmd(parser & p) {
return definition_cmd_core(p, Abbreviation, false, false);
return definition_cmd_core(p, Abbreviation, false, false, false);
}
environment local_abbreviation_cmd(parser & p) {
return definition_cmd_core(p, LocalAbbreviation, true, false);
return definition_cmd_core(p, LocalAbbreviation, true, false, false);
}
static environment theorem_cmd(parser & p) {
return definition_cmd_core(p, Theorem, false, false);
return definition_cmd_core(p, Theorem, false, false, false);
}
static environment example_cmd(parser & p) {
definition_cmd_core(p, Example, false, false);
definition_cmd_core(p, Example, false, false, false);
return p.env();
}
static environment private_definition_cmd(parser & p) {
bool is_noncomputable = false;
if (p.curr_is_token(get_noncomputable_tk())) {
is_noncomputable = true;
p.next();
}
def_cmd_kind kind = Definition;
if (p.curr_is_token_or_id(get_definition_tk())) {
p.next();
@ -1114,7 +1131,7 @@ static environment private_definition_cmd(parser & p) {
} else {
throw parser_error("invalid 'private' definition/theorem, 'definition' or 'theorem' expected", p.pos());
}
return definition_cmd_core(p, kind, true, false);
return definition_cmd_core(p, kind, true, false, is_noncomputable);
}
static environment protected_definition_cmd(parser & p) {
if (p.curr_is_token_or_id(get_axiom_tk())) {
@ -1130,6 +1147,11 @@ static environment protected_definition_cmd(parser & p) {
p.next();
return variables_cmd_core(p, variable_kind::Constant, true);
} else {
bool is_noncomputable = false;
if (p.curr_is_token(get_noncomputable_tk())) {
is_noncomputable = true;
p.next();
}
def_cmd_kind kind = Definition;
if (p.curr_is_token_or_id(get_definition_tk())) {
p.next();
@ -1142,9 +1164,33 @@ static environment protected_definition_cmd(parser & p) {
} else {
throw parser_error("invalid 'protected' definition/theorem, 'definition' or 'theorem' expected", p.pos());
}
return definition_cmd_core(p, kind, false, true);
return definition_cmd_core(p, kind, false, true, is_noncomputable);
}
}
static environment noncomputable_cmd(parser & p) {
bool is_private = false;
bool is_protected = false;
if (p.curr_is_token(get_private_tk())) {
is_private = true;
p.next();
} else if (p.curr_is_token(get_protected_tk())) {
is_protected = true;
p.next();
}
def_cmd_kind kind = Definition;
if (p.curr_is_token_or_id(get_definition_tk())) {
p.next();
} else if (p.curr_is_token_or_id(get_abbreviation_tk())) {
p.next();
kind = Abbreviation;
} else if (p.curr_is_token_or_id(get_theorem_tk())) {
p.next();
kind = Theorem;
} else {
throw parser_error("invalid 'noncomputable' definition/theorem, 'definition' or 'theorem' expected", p.pos());
}
return definition_cmd_core(p, kind, is_private, is_protected, true);
}
static environment include_cmd_core(parser & p, bool include) {
if (!p.curr_is_identifier())
@ -1222,6 +1268,7 @@ void register_decl_cmds(cmd_table & r) {
add_cmd(r, cmd_info("constants", "declare new constants (aka top-level variables)", constants_cmd));
add_cmd(r, cmd_info("axioms", "declare new axioms", axioms_cmd));
add_cmd(r, cmd_info("definition", "add new definition", definition_cmd));
add_cmd(r, cmd_info("noncomputable", "add new noncomputable definition", noncomputable_cmd));
add_cmd(r, cmd_info("example", "add new example", example_cmd));
add_cmd(r, cmd_info("private", "add new private definition/theorem", private_definition_cmd));
add_cmd(r, cmd_info("protected", "add new protected definition/theorem", protected_definition_cmd));

View file

@ -105,7 +105,7 @@ void init_token_table(token_table & t) {
char const * commands[] =
{"theorem", "axiom", "axioms", "variable", "protected", "private", "reveal",
"definition", "example", "coercion", "abbreviation",
"definition", "example", "coercion", "abbreviation", "noncomputable",
"variables", "parameter", "parameters", "constant", "constants", "[persistent]", "[visible]", "[instance]", "[trans-instance]",
"[none]", "[class]", "[coercion]", "[reducible]", "[irreducible]", "[semireducible]", "[quasireducible]",
"[simp]", "[congr]", "[parsing-only]", "[multiple-instances]", "[symm]", "[trans]", "[refl]", "[subst]", "[recursor",

View file

@ -88,6 +88,7 @@ static name const * g_begin_tk = nullptr;
static name const * g_begin_plus_tk = nullptr;
static name const * g_end_tk = nullptr;
static name const * g_private_tk = nullptr;
static name const * g_protected_tk = nullptr;
static name const * g_definition_tk = nullptr;
static name const * g_theorem_tk = nullptr;
static name const * g_abbreviation_tk = nullptr;
@ -151,6 +152,7 @@ static name const * g_trust_tk = nullptr;
static name const * g_metaclasses_tk = nullptr;
static name const * g_inductive_tk = nullptr;
static name const * g_this_tk = nullptr;
static name const * g_noncomputable_tk = nullptr;
void initialize_tokens() {
g_period_tk = new name{"."};
g_placeholder_tk = new name{"_"};
@ -237,6 +239,7 @@ void initialize_tokens() {
g_begin_plus_tk = new name{"begin+"};
g_end_tk = new name{"end"};
g_private_tk = new name{"private"};
g_protected_tk = new name{"protected"};
g_definition_tk = new name{"definition"};
g_theorem_tk = new name{"theorem"};
g_abbreviation_tk = new name{"abbreviation"};
@ -300,6 +303,7 @@ void initialize_tokens() {
g_metaclasses_tk = new name{"metaclasses"};
g_inductive_tk = new name{"inductive"};
g_this_tk = new name{"this"};
g_noncomputable_tk = new name{"noncomputable"};
}
void finalize_tokens() {
delete g_period_tk;
@ -387,6 +391,7 @@ void finalize_tokens() {
delete g_begin_plus_tk;
delete g_end_tk;
delete g_private_tk;
delete g_protected_tk;
delete g_definition_tk;
delete g_theorem_tk;
delete g_abbreviation_tk;
@ -450,6 +455,7 @@ void finalize_tokens() {
delete g_metaclasses_tk;
delete g_inductive_tk;
delete g_this_tk;
delete g_noncomputable_tk;
}
name const & get_period_tk() { return *g_period_tk; }
name const & get_placeholder_tk() { return *g_placeholder_tk; }
@ -536,6 +542,7 @@ name const & get_begin_tk() { return *g_begin_tk; }
name const & get_begin_plus_tk() { return *g_begin_plus_tk; }
name const & get_end_tk() { return *g_end_tk; }
name const & get_private_tk() { return *g_private_tk; }
name const & get_protected_tk() { return *g_protected_tk; }
name const & get_definition_tk() { return *g_definition_tk; }
name const & get_theorem_tk() { return *g_theorem_tk; }
name const & get_abbreviation_tk() { return *g_abbreviation_tk; }
@ -599,4 +606,5 @@ name const & get_trust_tk() { return *g_trust_tk; }
name const & get_metaclasses_tk() { return *g_metaclasses_tk; }
name const & get_inductive_tk() { return *g_inductive_tk; }
name const & get_this_tk() { return *g_this_tk; }
name const & get_noncomputable_tk() { return *g_noncomputable_tk; }
}

View file

@ -90,6 +90,7 @@ name const & get_begin_tk();
name const & get_begin_plus_tk();
name const & get_end_tk();
name const & get_private_tk();
name const & get_protected_tk();
name const & get_definition_tk();
name const & get_theorem_tk();
name const & get_abbreviation_tk();
@ -153,4 +154,5 @@ name const & get_trust_tk();
name const & get_metaclasses_tk();
name const & get_inductive_tk();
name const & get_this_tk();
name const & get_noncomputable_tk();
}

View file

@ -83,6 +83,7 @@ begin begin
begin_plus begin+
end end
private private
protected protected
definition definition
theorem theorem
abbreviation abbreviation
@ -146,3 +147,4 @@ trust trust
metaclasses metaclasses
inductive inductive
this this
noncomputable noncomputable

View file

@ -23,6 +23,7 @@ Author: Leonardo de Moura
#include "kernel/quotient/quotient.h"
#include "kernel/hits/hits.h"
#include "library/module.h"
#include "library/noncomputable.h"
#include "library/sorry.h"
#include "library/kernel_serializer.h"
#include "library/unfold_macros.h"
@ -219,11 +220,15 @@ static environment export_decl(environment const & env, declaration const & d) {
environment add(environment const & env, certified_declaration const & d) {
environment new_env = env.add(d);
declaration _d = d.get_declaration();
if (!check_computable(new_env, _d.get_name()))
new_env = mark_noncomputable(new_env, _d.get_name());
return export_decl(update_module_defs(new_env, _d), _d);
}
environment add(environment const & env, declaration const & d) {
environment new_env = env.add(d);
if (!check_computable(new_env, d.get_name()))
new_env = mark_noncomputable(new_env, d.get_name());
return export_decl(update_module_defs(new_env, d), d);
}

View file

@ -62,6 +62,11 @@ bool is_noncomputable(environment const & env, name const & n) {
return is_noncomputable(tc, ext, n);
}
bool is_marked_noncomputable(environment const & env, name const & n) {
auto ext = get_extension(env);
return ext.m_noncomputable.contains(n);
}
environment mark_noncomputable(environment const & env, name const & n) {
auto ext = get_extension(env);
ext.m_noncomputable.insert(n);
@ -69,23 +74,29 @@ environment mark_noncomputable(environment const & env, name const & n) {
return module::add(new_env, *g_key, [=](environment const &, serializer & s) { s << n; });
}
void validate_computable_definition(environment const & env, name const & n) {
optional<name> get_noncomputable_reason(environment const & env, name const & n) {
if (!is_standard(env))
return; // do nothing if it is not a standard kernel
type_checker tc(env);
return optional<name>(); // do nothing if it is not a standard kernel
declaration const & d = env.get(n);
if (!d.is_definition())
return;
return optional<name>();
type_checker tc(env);
if (tc.is_prop(d.get_type()).first)
return; // definition is a proposition, then do nothing
return optional<name>(); // definition is a proposition, then do nothing
expr const & v = d.get_value();
auto ext = get_extension(env);
optional<name> r;
for_each(v, [&](expr const & e, unsigned) {
if (is_constant(e) && is_noncomputable(tc, ext, const_name(e))) {
throw exception(sstream() << "definition '" << n << "' is noncomputable, it depends on '" << const_name(e) << "'");
r = const_name(e);
}
return true;
});
return r;
}
bool check_computable(environment const & env, name const & n) {
return !get_noncomputable_reason(env, n);
}
void initialize_noncomputable() {

View file

@ -8,12 +8,14 @@ Author: Leonardo de Moura
#include "kernel/environment.h"
namespace lean {
bool is_marked_noncomputable(environment const & env, name const & n);
/** \brief Return true if \c n is noncomputable */
bool is_noncomputable(environment const & env, name const & n);
/** \brief Mark \c n as noncomputable */
environment mark_noncomputable(environment const & env, name const & n);
/** \brief Throw an exception if \c n depends on a noncomputable definition */
void validate_computable_definition(environment const & env, name const & n);
/** \brief In standard mode, check if definitions that are not propositions can compute */
bool check_computable(environment const & env, name const & n);
optional<name> get_noncomputable_reason(environment const & env, name const & n);
void initialize_noncomputable();
void finalize_noncomputable();
}

View file

@ -1,6 +1,6 @@
import logic logic.axioms.hilbert
open inhabited nonempty
definition v1 : Prop := epsilon (λ x, true)
noncomputable definition v1 : Prop := epsilon (λ x, true)
inductive Empty : Type
definition v2 : Empty := epsilon (λ x, true)
noncomputable definition v2 : Empty := epsilon (λ x, true)

View file

@ -1,3 +1,3 @@
empty.lean:6:25: error: failed to synthesize placeholder
empty.lean:6:39: error: failed to synthesize placeholder
⊢ nonempty Empty

View file

@ -30,7 +30,7 @@ end foo
open nat
definition bla : nat :=
noncomputable definition bla : nat :=
foo.tst1 0 0 + foo.tst2 0 0 + foo.tst3 nat 1 1
check foo.tst1

8
tests/lean/noncomp.lean Normal file
View file

@ -0,0 +1,8 @@
open nat
axiom n : nat
definition f (x : nat) := -- Error this is not computable
x + n
noncomputable definition f (x : nat) := x + n

View file

@ -0,0 +1 @@
noncomp.lean:5:0: error: definition 'f' is noncomputable, it depends on 'n'

View file

@ -12,5 +12,5 @@ attribute path_fibrant [instance]
axiom imp_fibrant {A : Type'} {B : Type'} [C1 : fibrant A] [C2 : fibrant B] : fibrant (A → B)
attribute imp_fibrant [instance]
definition test {A : Type} [fA : fibrant A] {x y : A} :
noncomputable definition test {A : Type} [fA : fibrant A] {x y : A} :
Π (z : A), y ≈ z → fibrant (x ≈ y → x ≈ z) := _

View file

@ -13,20 +13,20 @@ attribute path_fibrant [instance]
notation a ≈ b := path a b
definition test {A : Type} [fA : fibrant A] {x y : A} :
noncomputable definition test {A : Type} [fA : fibrant A] {x y : A} :
Π (z : A), y ≈ z → fibrant (x ≈ y → x ≈ z) := take z p, _
definition test2 {A : Type} [fA : fibrant A] {x y : A} :
noncomputable definition test2 {A : Type} [fA : fibrant A] {x y : A} :
Π (z : A), y ≈ z → fibrant (x ≈ y → x ≈ z) := _
definition test3 {A : Type} [fA : fibrant A] {x y : A} :
noncomputable definition test3 {A : Type} [fA : fibrant A] {x y : A} :
Π (z : A), y ≈ z → fibrant (x ≈ z) := _
definition test4 {A : Type} [fA : fibrant A] {x y z : A} :
noncomputable definition test4 {A : Type} [fA : fibrant A] {x y z : A} :
fibrant (x ≈ y → x ≈ z) := _
axiom imp_fibrant {A : Type} {B : Type} [C1 : fibrant A] [C2 : fibrant B] : fibrant (A → B)
attribute imp_fibrant [instance]
definition test5 {A : Type} [fA : fibrant A] {x y : A} :
noncomputable definition test5 {A : Type} [fA : fibrant A] {x y : A} :
Π (z : A), y ≈ z → fibrant (x ≈ y → x ≈ z) := _

View file

@ -14,7 +14,7 @@ theorem inr_inhabited (A : Type) {B : Type} (H : inhabited B) : inhabited (sum A
notation `(` h `|` r:(foldl `|` (e r, tactic.or_else r e) h) `)` := r
infixl `;`:15 := tactic.and_then
reveal inl_inhabited inr_inhabited
definition my_tac := fixpoint (λ t, ( apply @inl_inhabited; t
| apply @inr_inhabited; t
| apply @num.is_inhabited

View file

@ -16,7 +16,7 @@ infixl `..`:10 := append
notation `(` h `|` r:(foldl `|` (e r, tactic.or_else r e) h) `)` := r
infixl `;`:15 := tactic.and_then
reveal inl_inhabited inr_inhabited
definition my_tac := repeat (trace "iteration"; state;
( apply @inl_inhabited; trace "used inl"
.. apply @inr_inhabited; trace "used inr"

View file

@ -391,10 +391,12 @@ theorem triple_induction
C x y z w p q r :=
path.rec_on p (take z q, path.rec_on q (take w r, path.rec_on r H)) z q w r
reveal path.double_induction2
-- try this again
definition concat_pV_p_new {A : Type} {x y z : A} (p : x ≈ z) (q : y ≈ z) : (p ⬝ q**) ⬝ q ≈ p :=
double_induction2 p q idp
reveal path.triple_induction
definition transport_p_pp {A : Type} (P : A → Type)
{x y z w : A} (p : x ≈ y) (q : y ≈ z) (r : z ≈ w) (u : P x) :
ap (λe, e # u) (concat_p_pp p q r) ⬝ (transport_pp P (p ⬝ q) r u) ⬝