diff --git a/src/builtin/kernel.lean b/src/builtin/kernel.lean index 48a19db34..78aa51c00 100644 --- a/src/builtin/kernel.lean +++ b/src/builtin/kernel.lean @@ -809,7 +809,7 @@ theorem inhabited_fun (A : (Type U)) {B : (Type U)} (H : inhabited B) : inhabite theorem exists_to_eps {A : (Type U)} {P : A → Bool} (H : ∃ x, P x) : P (ε (inhabited_ex_intro H) P) := obtain (w : A) (Hw : P w), from H, - eps_ax (inhabited_ex_intro H) w Hw + @eps_ax _ (inhabited_ex_intro H) P w Hw theorem axiom_of_choice {A : (Type U)} {B : A → (Type U)} {R : ∀ x : A, B x → Bool} (H : ∀ x, ∃ y, R x y) : ∃ f, ∀ x, R x (f x) := exists_intro @@ -819,7 +819,7 @@ theorem axiom_of_choice {A : (Type U)} {B : A → (Type U)} {R : ∀ x : A, B x theorem skolem_th {A : (Type U)} {B : A → (Type U)} {P : ∀ x : A, B x → Bool} : (∀ x, ∃ y, P x y) ↔ ∃ f, (∀ x, P x (f x)) := iff_intro - (λ H : (∀ x, ∃ y, P x y), axiom_of_choice H) + (λ H : (∀ x, ∃ y, P x y), @axiom_of_choice _ _ P H) (λ H : (∃ f, (∀ x, P x (f x))), take x, obtain (fw : ∀ x, B x) (Hw : ∀ x, P x (fw x)), from H, exists_intro (fw x) (Hw x)) diff --git a/src/builtin/obj/Nat.olean b/src/builtin/obj/Nat.olean index f9391e528..8c23f25d8 100644 Binary files a/src/builtin/obj/Nat.olean and b/src/builtin/obj/Nat.olean differ diff --git a/src/builtin/obj/kernel.olean b/src/builtin/obj/kernel.olean index 1f25e1e0c..8e5103111 100644 Binary files a/src/builtin/obj/kernel.olean and b/src/builtin/obj/kernel.olean differ diff --git a/src/builtin/obj/num.olean b/src/builtin/obj/num.olean index d904b9b9a..7c21f7b23 100644 Binary files a/src/builtin/obj/num.olean and b/src/builtin/obj/num.olean differ diff --git a/src/builtin/obj/optional.olean b/src/builtin/obj/optional.olean index d559e0fe1..5d99ce1dd 100644 Binary files a/src/builtin/obj/optional.olean and b/src/builtin/obj/optional.olean differ diff --git a/src/builtin/obj/subtype.olean b/src/builtin/obj/subtype.olean index bc85d570b..fbadac036 100644 Binary files a/src/builtin/obj/subtype.olean and b/src/builtin/obj/subtype.olean differ diff --git a/src/builtin/obj/sum.olean b/src/builtin/obj/sum.olean index adc566891..31758cfe8 100644 Binary files a/src/builtin/obj/sum.olean and b/src/builtin/obj/sum.olean differ diff --git a/src/library/elaborator/elaborator.cpp b/src/library/elaborator/elaborator.cpp index dc3a75633..a37bf8ae5 100644 --- a/src/library/elaborator/elaborator.cpp +++ b/src/library/elaborator/elaborator.cpp @@ -895,7 +895,8 @@ class elaborator::imp { The constraint is solved by assigning ?m to (fun x1 ... xn, c). The arguments may contain duplicate occurrences of variables. - + The following case is INCORRECT, and restricts the set of solutions that can be found. + IT WAS DISABLED. 2) a constraint of the form ctx |- (?m x1 ... xn) ≈ (f x1 ... xn) where f does not contain x1 ... xn, and x1 ... xn are distinct The constraint is solved by assigning ?m to f OR ?m to (fun x1 ... xn, f x1 ... xn) @@ -922,6 +923,14 @@ class elaborator::imp { push_new_eq_constraint(ctx, m, s, new_jst); return true; } else { + return false; + #if 0 + // This code was disabled because it removes valid solutions. + // For example, when this piece of code is enabled, we can't + // solve + // theorem or_imp (p q : Bool) : (p ∨ q) ↔ (¬ p → q) + // := subst (symm (imp_or (¬ p) q)) (not_not_eq p) + // TODO(Leo): delete this piece of code lean_assert(eq_app); expr f = arg(b, 0); std::unique_ptr new_cs(new generic_case_split(c, m_state)); @@ -946,6 +955,7 @@ class elaborator::imp { lean_assert(r); m_case_splits.push_back(std::move(new_cs)); return r; + #endif } } else { return false; diff --git a/tests/lean/elab4.lean.expected.out b/tests/lean/elab4.lean.expected.out index 91aace6d2..144ace9e8 100644 --- a/tests/lean/elab4.lean.expected.out +++ b/tests/lean/elab4.lean.expected.out @@ -15,6 +15,6 @@ variable R {A A' : Type} {B' : A' → Type} (H : @eq Type (∀ x : A, B x) (∀ x : A', B' x)) (a : A) : - @eq Type (B a) (B' (@C A A' (@D A A' B B' H) a)) + @eq Type (B a) (B' (@C A A' (@D A A' (λ x : A, B x) (λ x : A', B' x) H) a)) theorem R2 (A1 A2 B1 B2 : Type) (H : @eq Type (A1 → B1) (A2 → B2)) (a : A1) : @eq Type B1 B2 := @R A1 A2 (λ x : A1, B1) (λ x : A2, B2) H a diff --git a/tests/lean/elab5.lean.expected.out b/tests/lean/elab5.lean.expected.out index 91aace6d2..144ace9e8 100644 --- a/tests/lean/elab5.lean.expected.out +++ b/tests/lean/elab5.lean.expected.out @@ -15,6 +15,6 @@ variable R {A A' : Type} {B' : A' → Type} (H : @eq Type (∀ x : A, B x) (∀ x : A', B' x)) (a : A) : - @eq Type (B a) (B' (@C A A' (@D A A' B B' H) a)) + @eq Type (B a) (B' (@C A A' (@D A A' (λ x : A, B x) (λ x : A', B' x) H) a)) theorem R2 (A1 A2 B1 B2 : Type) (H : @eq Type (A1 → B1) (A2 → B2)) (a : A1) : @eq Type B1 B2 := @R A1 A2 (λ x : A1, B1) (λ x : A2, B2) H a diff --git a/tests/lean/elab7.lean.expected.out b/tests/lean/elab7.lean.expected.out index 082369d5d..c3d3843ee 100644 --- a/tests/lean/elab7.lean.expected.out +++ b/tests/lean/elab7.lean.expected.out @@ -9,8 +9,9 @@ Assumed: H Assumed: a eta (F2 a) : (λ x : B, F2 a x) = F2 a -funext (λ a : A, trans (symm (eta (F1 a))) (trans (funext (λ b : B, H a b)) (eta (F2 a)))) : F1 = F2 -funext (λ a : A, funext (λ b : B, H a b)) : F1 = F2 +funext (λ a : A, trans (symm (eta (F1 a))) (trans (funext (λ b : B, H a b)) (eta (F2 a)))) : + (λ x : A, F1 x) = (λ x : A, F2 x) +funext (λ a : A, funext (λ b : B, H a b)) : (λ (x : A) (x::1 : B), F1 x x::1) = (λ (x : A) (x::1 : B), F2 x x::1) Proved: T1 Proved: T2 Proved: T3 diff --git a/tests/lean/j5.lean b/tests/lean/j5.lean new file mode 100644 index 000000000..2c929319f --- /dev/null +++ b/tests/lean/j5.lean @@ -0,0 +1,3 @@ +theorem or_imp (p q : Bool) : (p ∨ q) ↔ (¬ p → q) +:= subst (symm (imp_or (¬ p) q)) (not_not_eq p) + diff --git a/tests/lean/j5.lean.expected.out b/tests/lean/j5.lean.expected.out new file mode 100644 index 000000000..6ddc0d932 --- /dev/null +++ b/tests/lean/j5.lean.expected.out @@ -0,0 +1,3 @@ + Set: pp::colors + Set: pp::unicode + Proved: or_imp diff --git a/tests/lean/j6.lean b/tests/lean/j6.lean new file mode 100644 index 000000000..437a60c80 --- /dev/null +++ b/tests/lean/j6.lean @@ -0,0 +1,8 @@ +theorem symm_iff (p q : Bool) (H : p ↔ q) : (q ↔ p) +:= symm H + +theorem or_imp (p q : Bool) : (p ∨ q) ↔ (¬ p → q) +:= let H1 := symm_iff _ _ (imp_or (¬ p) q) in + let H2 := not_not_eq p in + let H3 := subst H1 H2 in + H3 diff --git a/tests/lean/j6.lean.expected.out b/tests/lean/j6.lean.expected.out new file mode 100644 index 000000000..abe3c1281 --- /dev/null +++ b/tests/lean/j6.lean.expected.out @@ -0,0 +1,4 @@ + Set: pp::colors + Set: pp::unicode + Proved: symm_iff + Proved: or_imp