fix(lean/library/standard): fix tests, more cleanup
This commit is contained in:
parent
148d475421
commit
6264fb52d6
36 changed files with 120 additions and 154 deletions
|
@ -3,6 +3,7 @@
|
|||
-- Author: Leonardo de Moura
|
||||
|
||||
import logic.connectives.basic logic.classes.decidable logic.classes.inhabited
|
||||
|
||||
using eq_ops decidable
|
||||
|
||||
inductive bool : Type :=
|
||||
|
@ -14,7 +15,7 @@ namespace bool
|
|||
theorem induction_on {p : bool → Prop} (b : bool) (H0 : p ff) (H1 : p tt) : p b :=
|
||||
bool_rec H0 H1 b
|
||||
|
||||
theorem inhabited_bool [instance] : inhabited bool :=
|
||||
theorem bool_inhabited [instance] : inhabited bool :=
|
||||
inhabited_mk ff
|
||||
|
||||
definition cond {A : Type} (b : bool) (t e : A) :=
|
||||
|
|
|
@ -23,7 +23,7 @@ inductive num : Type :=
|
|||
theorem inhabited_pos_num [instance] : inhabited pos_num :=
|
||||
inhabited_mk one
|
||||
|
||||
theorem inhabited_num [instance] : inhabited num :=
|
||||
theorem num_inhabited [instance] : inhabited num :=
|
||||
inhabited_mk zero
|
||||
|
||||
end num
|
||||
|
|
|
@ -1,20 +1,22 @@
|
|||
----------------------------------------------------------------------------------------------------
|
||||
-- Copyright (c) 2014 Microsoft Corporation. All rights reserved.
|
||||
-- Released under Apache 2.0 license as described in the file LICENSE.
|
||||
-- Author: Leonardo de Moura
|
||||
----------------------------------------------------------------------------------------------------
|
||||
|
||||
import logic.connectives.basic logic.connectives.eq logic.classes.inhabited logic.classes.decidable
|
||||
using eq_ops decidable
|
||||
|
||||
namespace option
|
||||
|
||||
inductive option (A : Type) : Type :=
|
||||
| none {} : option A
|
||||
| some : A → option A
|
||||
|
||||
theorem induction_on {A : Type} {p : option A → Prop} (o : option A) (H1 : p none) (H2 : ∀a, p (some a)) : p o :=
|
||||
theorem induction_on {A : Type} {p : option A → Prop} (o : option A)
|
||||
(H1 : p none) (H2 : ∀a, p (some a)) : p o :=
|
||||
option_rec H1 H2 o
|
||||
|
||||
definition rec_on {A : Type} {C : option A → Type} (o : option A) (H1 : C none) (H2 : ∀a, C (some a)) : C o :=
|
||||
definition rec_on {A : Type} {C : option A → Type} (o : option A)
|
||||
(H1 : C none) (H2 : ∀a, C (some a)) : C o :=
|
||||
option_rec H1 H2 o
|
||||
|
||||
definition is_none {A : Type} (o : option A) : Prop :=
|
||||
|
@ -34,10 +36,11 @@ assume H : none = some a, absurd
|
|||
theorem some_inj {A : Type} {a₁ a₂ : A} (H : some a₁ = some a₂) : a₁ = a₂ :=
|
||||
congr_arg (option_rec a₁ (λ a, a)) H
|
||||
|
||||
theorem inhabited_option [instance] (A : Type) : inhabited (option A) :=
|
||||
theorem option_inhabited [instance] (A : Type) : inhabited (option A) :=
|
||||
inhabited_mk none
|
||||
|
||||
theorem decidable_eq [instance] {A : Type} {H : ∀a₁ a₂ : A, decidable (a₁ = a₂)} (o₁ o₂ : option A) : decidable (o₁ = o₂) :=
|
||||
theorem decidable_eq [instance] {A : Type} {H : ∀a₁ a₂ : A, decidable (a₁ = a₂)}
|
||||
(o₁ o₂ : option A) : decidable (o₁ = o₂) :=
|
||||
rec_on o₁
|
||||
(rec_on o₂ (inl (refl _)) (take a₂, (inr (none_ne_some a₂))))
|
||||
(take a₁ : A, rec_on o₂
|
||||
|
@ -45,4 +48,5 @@ rec_on o₁
|
|||
(take a₂ : A, decidable.rec_on (H a₁ a₂)
|
||||
(assume Heq : a₁ = a₂, inl (Heq ▸ refl _))
|
||||
(assume Hne : a₁ ≠ a₂, inr (assume Hn : some a₁ = some a₂, absurd (some_inj Hn) Hne))))
|
||||
|
||||
end option
|
||||
|
|
|
@ -7,13 +7,14 @@ import logic.connectives.instances
|
|||
import .aux
|
||||
|
||||
using relation prod inhabited nonempty tactic eq_ops
|
||||
using subtype relation.iff_ops
|
||||
|
||||
|
||||
-- Theory data.quotient
|
||||
-- ====================
|
||||
|
||||
namespace quotient
|
||||
|
||||
using subtype
|
||||
|
||||
-- definition and basics
|
||||
-- ---------------------
|
||||
|
@ -29,34 +30,6 @@ theorem intro {A B : Type} {R : A → A → Prop} {abs : A → B} {rep : B → A
|
|||
(H3 : ∀r s, R r s ↔ (R r r ∧ R s s ∧ abs r = abs s)) : is_quotient R abs rep :=
|
||||
and_intro H1 (and_intro H2 H3)
|
||||
|
||||
-- theorem intro_refl {A B : Type} {R : A → A → Prop} {abs : A → B} {rep : B → A}
|
||||
-- (H1 : reflexive R) (H2 : ∀b, abs (rep b) = b)
|
||||
-- (H3 : ∀r s, R r s ↔ abs r = abs s) : is_quotient R abs rep :=
|
||||
-- intro
|
||||
-- H2
|
||||
-- (take b, H1 (rep b))
|
||||
-- (take r s,
|
||||
-- have H4 : R r s ↔ R s s ∧ abs r = abs s,
|
||||
-- from
|
||||
-- gensubst.subst (relation.operations.symm (and_absorb_left _ (H1 s))) (H3 r s),
|
||||
-- gensubst.subst (relation.operations.symm (and_absorb_left _ (H1 r))) H4)
|
||||
|
||||
-- these work now, but the above still does not
|
||||
-- theorem test (a b c : Prop) (P : Prop → Prop) (H1 : a ↔ b) (H2 : c ∧ a) : c ∧ b :=
|
||||
-- gensubst.subst H1 H2
|
||||
|
||||
-- theorem test2 {A : Type} {R : A → A → Prop} (Q : Prop) (r s : A)
|
||||
-- (H3 : R r s ↔ Q) (H1 : R s s) : Q ↔ (R s s ∧ Q) :=
|
||||
-- relation.operations.symm (and_absorb_left Q H1)
|
||||
|
||||
-- theorem test3 {A : Type} {R : A → A → Prop} (Q : Prop) (r s : A)
|
||||
-- (H3 : R r s ↔ Q) (H1 : R s s) : R r s ↔ (R s s ∧ Q) :=
|
||||
-- gensubst.subst (test2 Q r s H3 H1) H3
|
||||
|
||||
-- theorem test4 {A : Type} {R : A → A → Prop} (Q : Prop) (r s : A)
|
||||
-- (H3 : R r s ↔ Q) (H1 : R s s) : R r s ↔ (R s s ∧ Q) :=
|
||||
-- gensubst.subst (relation.operations.symm (and_absorb_left Q H1)) H3
|
||||
|
||||
theorem and_absorb_left {a : Prop} (b : Prop) (Ha : a) : a ∧ b ↔ b :=
|
||||
iff_intro (assume Hab, and_elim_right Hab) (assume Hb, and_intro Ha Hb)
|
||||
|
||||
|
@ -69,20 +42,8 @@ intro
|
|||
(take r s,
|
||||
have H4 : R r s ↔ R s s ∧ abs r = abs s,
|
||||
from
|
||||
subst_iff (iff_symm (and_absorb_left _ (H1 s))) (H3 r s),
|
||||
subst_iff (iff_symm (and_absorb_left _ (H1 r))) H4)
|
||||
|
||||
-- theorem intro_refl {A B : Type} {R : A → A → Prop} {abs : A → B} {rep : B → A}
|
||||
-- (H1 : reflexive R) (H2 : ∀b, abs (rep b) = b)
|
||||
-- (H3 : ∀r s, R r s ↔ abs r = abs s) : is_quotient R abs rep :=
|
||||
-- intro
|
||||
-- H2
|
||||
-- (take b, H1 (rep b))
|
||||
-- (take r s,
|
||||
-- have H4 : R r s ↔ R s s ∧ abs r = abs s,
|
||||
-- from
|
||||
-- substi (iff_symm (and_absorb_left _ (H1 s))) (H3 r s),
|
||||
-- substi (iff_symm (and_absorb_left _ (H1 r))) H4)
|
||||
subst (symm (and_absorb_left _ (H1 s))) (H3 r s),
|
||||
subst (symm (and_absorb_left _ (H1 r))) H4)
|
||||
|
||||
theorem abs_rep {A B : Type} {R : A → A → Prop} {abs : A → B} {rep : B → A}
|
||||
(Q : is_quotient R abs rep) (b : B) : abs (rep b) = b :=
|
||||
|
@ -147,6 +108,7 @@ have Hc : R c c, from refl_right Q Hbc,
|
|||
have Hac : abs a = abs c, from trans (eq_abs Q Hab) (eq_abs Q Hbc),
|
||||
R_intro Q Ha Hc Hac
|
||||
|
||||
|
||||
-- recursion
|
||||
-- ---------
|
||||
|
||||
|
@ -299,28 +261,28 @@ calc
|
|||
-- ------------------------------------------
|
||||
|
||||
theorem representative_map_idempotent {A : Type} {R : A → A → Prop} {f : A → A}
|
||||
(H1 : ∀a, R a (f a)) (H2 : ∀a b, R a b ↔ R a a ∧ R b b ∧ f a = f b) (a : A)
|
||||
: f (f a) = f a :=
|
||||
(H1 : ∀a, R a (f a)) (H2 : ∀a b, R a b ↔ R a a ∧ R b b ∧ f a = f b) (a : A) :
|
||||
f (f a) = f a :=
|
||||
symm (and_elim_right (and_elim_right (iff_elim_left (H2 a (f a)) (H1 a))))
|
||||
|
||||
theorem representative_map_idempotent_equiv {A : Type} {R : A → A → Prop} {f : A → A}
|
||||
(H1 : ∀a, R a (f a)) (H2 : ∀a b, R a b → f a = f b) (a : A)
|
||||
: f (f a) = f a :=
|
||||
(H1 : ∀a, R a (f a)) (H2 : ∀a b, R a b → f a = f b) (a : A) :
|
||||
f (f a) = f a :=
|
||||
symm (H2 a (f a) (H1 a))
|
||||
|
||||
theorem representative_map_refl_rep {A : Type} {R : A → A → Prop} {f : A → A}
|
||||
(H1 : ∀a, R a (f a)) (H2 : ∀a b, R a b ↔ R a a ∧ R b b ∧ f a = f b) (a : A)
|
||||
: R (f a) (f a) :=
|
||||
(H1 : ∀a, R a (f a)) (H2 : ∀a b, R a b ↔ R a a ∧ R b b ∧ f a = f b) (a : A) :
|
||||
R (f a) (f a) :=
|
||||
subst (representative_map_idempotent H1 H2 a) (H1 (f a))
|
||||
|
||||
theorem representative_map_image_fix {A : Type} {R : A → A → Prop} {f : A → A}
|
||||
(H1 : ∀a, R a (f a)) (H2 : ∀a a', R a a' ↔ R a a ∧ R a' a' ∧ f a = f a') (b : image f)
|
||||
: f (elt_of b) = elt_of b :=
|
||||
(H1 : ∀a, R a (f a)) (H2 : ∀a a', R a a' ↔ R a a ∧ R a' a' ∧ f a = f a') (b : image f) :
|
||||
f (elt_of b) = elt_of b :=
|
||||
idempotent_image_fix (representative_map_idempotent H1 H2) b
|
||||
|
||||
theorem representative_map_to_quotient {A : Type} {R : A → A → Prop} {f : A → A}
|
||||
(H1 : ∀a, R a (f a)) (H2 : ∀a a', R a a' ↔ R a a ∧ R a' a' ∧ f a = f a')
|
||||
: is_quotient _ (fun_image f) elt_of :=
|
||||
(H1 : ∀a, R a (f a)) (H2 : ∀a a', R a a' ↔ R a a ∧ R a' a' ∧ f a = f a') :
|
||||
is_quotient _ (fun_image f) elt_of :=
|
||||
let abs [inline] := fun_image f in
|
||||
intro
|
||||
(take u : image f,
|
||||
|
@ -337,30 +299,31 @@ intro
|
|||
obtain (a : A) (Ha : f a = elt_of u), from image_elt_of u,
|
||||
subst Ha (@representative_map_refl_rep A R f H1 H2 a))
|
||||
(take a a',
|
||||
subst_iff (fun_image_eq f a a') (H2 a a'))
|
||||
subst (fun_image_eq f a a') (H2 a a'))
|
||||
|
||||
theorem representative_map_equiv_inj {A : Type} {R : A → A → Prop}
|
||||
(equiv : is_equivalence R) {f : A → A} (H1 : ∀a, R a (f a)) (H2 : ∀a b, R a b → f a = f b)
|
||||
{a b : A} (H3 : f a = f b) : R a b :=
|
||||
have symmR : symmetric R, from is_symmetric.infer R,
|
||||
have transR : transitive R, from is_transitive.infer R,
|
||||
(equiv : is_equivalence R) {f : A → A}
|
||||
(H1 : ∀a, R a (f a)) (H2 : ∀a b, R a b → f a = f b) {a b : A} (H3 : f a = f b) :
|
||||
R a b :=
|
||||
have symmR : symmetric R, from rel_symm R,
|
||||
have transR : transitive R, from rel_trans R,
|
||||
show R a b, from
|
||||
have H2 : R a (f b), from subst H3 (H1 a),
|
||||
have H3 : R (f b) b, from symmR (H1 b),
|
||||
transR H2 H3
|
||||
|
||||
theorem representative_map_to_quotient_equiv {A : Type} {R : A → A → Prop}
|
||||
(equiv : is_equivalence R) {f : A → A} (H1 : ∀a, R a (f a)) (H2 : ∀a b, R a b → f a = f b)
|
||||
: @is_quotient A (image f) R (fun_image f) elt_of :=
|
||||
(equiv : is_equivalence R) {f : A → A} (H1 : ∀a, R a (f a)) (H2 : ∀a b, R a b → f a = f b) :
|
||||
@is_quotient A (image f) R (fun_image f) elt_of :=
|
||||
representative_map_to_quotient
|
||||
H1
|
||||
(take a b,
|
||||
have reflR : reflexive R, from is_reflexive.infer R,
|
||||
have reflR : reflexive R, from rel_refl R,
|
||||
have H3 : f a = f b → R a b, from representative_map_equiv_inj equiv H1 H2,
|
||||
have H4 : R a b ↔ f a = f b, from iff_intro (H2 a b) H3,
|
||||
have H5 : R a b ↔ R b b ∧ f a = f b,
|
||||
from subst_iff (iff_symm (and_absorb_left _ (reflR b))) H4,
|
||||
subst_iff (iff_symm (and_absorb_left _ (reflR a))) H5)
|
||||
from subst (symm (and_absorb_left _ (reflR b))) H4,
|
||||
subst (symm (and_absorb_left _ (reflR a))) H5)
|
||||
|
||||
-- previously:
|
||||
-- opaque_hint (hiding fun_image rec is_quotient prelim_map)
|
||||
|
|
|
@ -15,10 +15,10 @@ inductive string : Type :=
|
|||
| empty : string
|
||||
| str : char → string → string
|
||||
|
||||
theorem inhabited_char [instance] : inhabited char :=
|
||||
theorem char_inhabited [instance] : inhabited char :=
|
||||
inhabited_mk (ascii ff ff ff ff ff ff ff ff)
|
||||
|
||||
theorem inhabited_string [instance] : inhabited string :=
|
||||
theorem string_inhabited [instance] : inhabited string :=
|
||||
inhabited_mk empty
|
||||
|
||||
end string
|
||||
|
|
|
@ -1,13 +1,13 @@
|
|||
----------------------------------------------------------------------------------------------------
|
||||
-- Copyright (c) 2014 Microsoft Corporation. All rights reserved.
|
||||
-- Released under Apache 2.0 license as described in the file LICENSE.
|
||||
-- Author: Leonardo de Moura
|
||||
----------------------------------------------------------------------------------------------------
|
||||
|
||||
import logic.classes.decidable logic.classes.inhabited
|
||||
|
||||
using decidable
|
||||
|
||||
namespace unit
|
||||
|
||||
inductive unit : Type :=
|
||||
| star : unit
|
||||
|
||||
|
@ -16,9 +16,10 @@ notation `⋆`:max := star
|
|||
theorem unit_eq (a b : unit) : a = b :=
|
||||
unit_rec (unit_rec (refl ⋆) b) a
|
||||
|
||||
theorem inhabited_unit [instance] : inhabited unit :=
|
||||
theorem unit_inhabited [instance] : inhabited unit :=
|
||||
inhabited_mk ⋆
|
||||
|
||||
theorem decidable_eq [instance] (a b : unit) : decidable (a = b) :=
|
||||
inl (unit_eq a b)
|
||||
|
||||
end unit
|
||||
|
|
|
@ -27,7 +27,7 @@ theorem nonempty_imp_inhabited {A : Type} (H : nonempty A) : inhabited A :=
|
|||
let u : {x : A | (∃x : A, true) → true} := strong_indefinite_description (λa, true) H in
|
||||
inhabited_mk (elt_of u)
|
||||
|
||||
theorem inhabited_exists {A : Type} {P : A → Prop} (H : ∃x, P x) : inhabited A :=
|
||||
theorem exists_imp_inhabited {A : Type} {P : A → Prop} (H : ∃x, P x) : inhabited A :=
|
||||
nonempty_imp_inhabited (obtain w Hw, from H, nonempty_intro w)
|
||||
|
||||
|
||||
|
|
|
@ -10,13 +10,13 @@ using decidable inhabited nonempty
|
|||
-- Excluded middle + Hilbert implies every proposition is decidable
|
||||
|
||||
-- First, we show that (decidable a) is inhabited for any 'a' using the excluded middle
|
||||
theorem inhabited_decidable [instance] (a : Prop) : inhabited (decidable a) :=
|
||||
theorem decidable_inhabited [instance] (a : Prop) : inhabited (decidable a) :=
|
||||
nonempty_imp_inhabited
|
||||
(or_elim (em a)
|
||||
(assume Ha, nonempty_intro (inl Ha))
|
||||
(assume Hna, nonempty_intro (inr Hna)))
|
||||
|
||||
-- Note that inhabited_decidable is marked as an instance, and it is silently used
|
||||
-- Note that decidable_inhabited is marked as an instance, and it is silently used
|
||||
-- for synthesizing the implicit argument in the following 'epsilon'
|
||||
theorem prop_decidable [instance] (a : Prop) : decidable a :=
|
||||
epsilon (λd, true)
|
||||
|
|
|
@ -12,10 +12,10 @@ namespace inhabited
|
|||
definition inhabited_destruct {A : Type} {B : Type} (H1 : inhabited A) (H2 : A → B) : B :=
|
||||
inhabited_rec H2 H1
|
||||
|
||||
definition inhabited_Prop [instance] : inhabited Prop :=
|
||||
definition Prop_inhabited [instance] : inhabited Prop :=
|
||||
inhabited_mk true
|
||||
|
||||
definition inhabited_fun [instance] (A : Type) {B : Type} (H : inhabited B) : inhabited (A → B) :=
|
||||
definition fun_inhabited [instance] (A : Type) {B : Type} (H : inhabited B) : inhabited (A → B) :=
|
||||
inhabited_destruct H (take b, inhabited_mk (λa, b))
|
||||
|
||||
definition default (A : Type) {H : inhabited A} : A := inhabited_destruct H (take a, a)
|
||||
|
|
|
@ -153,10 +153,12 @@ theorem iff_intro {a b : Prop} (H1 : a → b) (H2 : b → a) : a ↔ b := and_in
|
|||
|
||||
theorem iff_elim {a b c : Prop} (H1 : (a → b) → (b → a) → c) (H2 : a ↔ b) : c := and_rec H1 H2
|
||||
|
||||
theorem iff_elim_left {a b : Prop} (H : a ↔ b) : a → b :=
|
||||
theorem iff_elim_left ⦃a b : Prop⦄ (H : a ↔ b) : a → b :=
|
||||
iff_elim (assume H1 H2, H1) H
|
||||
|
||||
theorem iff_elim_right {a b : Prop} (H : a ↔ b) : b → a :=
|
||||
abbreviation iff_mp := iff_elim_left
|
||||
|
||||
theorem iff_elim_right ⦃a b : Prop⦄ (H : a ↔ b) : b → a :=
|
||||
iff_elim (assume H1 H2, H2) H
|
||||
|
||||
theorem iff_flip_sign {a b : Prop} (H1 : a ↔ b) : ¬a ↔ ¬b :=
|
||||
|
@ -167,12 +169,12 @@ iff_intro
|
|||
theorem iff_refl (a : Prop) : a ↔ a :=
|
||||
iff_intro (assume H, H) (assume H, H)
|
||||
|
||||
theorem iff_trans {a b c : Prop} (H1 : a ↔ b) (H2 : b ↔ c) : a ↔ c :=
|
||||
theorem iff_trans ⦃a b c : Prop⦄ (H1 : a ↔ b) (H2 : b ↔ c) : a ↔ c :=
|
||||
iff_intro
|
||||
(assume Ha, iff_elim_left H2 (iff_elim_left H1 Ha))
|
||||
(assume Hc, iff_elim_right H1 (iff_elim_right H2 Hc))
|
||||
|
||||
theorem iff_symm {a b : Prop} (H : a ↔ b) : b ↔ a :=
|
||||
theorem iff_symm ⦃a b : Prop⦄ (H : a ↔ b) : b ↔ a :=
|
||||
iff_intro
|
||||
(assume Hb, iff_elim_right H Hb)
|
||||
(assume Ha, iff_elim_left H Ha)
|
||||
|
|
|
@ -44,6 +44,6 @@ theorem test7 (T : Type) (a b c d : T) (H1 : a = b) (H2 : c = b) (H3 : c = d) :
|
|||
trans H1 (trans (symm H2) H3)
|
||||
|
||||
theorem test8 (a b c d : Prop) (H1 : a ↔ b) (H2 : c ↔ b) (H3 : c ↔ d) : a ↔ d :=
|
||||
trans iff H1 (trans iff (symm iff H2) H3)
|
||||
trans H1 (trans (symm H2) H3)
|
||||
|
||||
end
|
||||
|
|
|
@ -107,30 +107,34 @@ relation.is_transitive_mk (@iff_trans)
|
|||
theorem mp_like_iff [instance] (a b : Prop) (H : a ↔ b) : relation.mp_like H :=
|
||||
relation.mp_like_mk (iff_elim_left H)
|
||||
|
||||
|
||||
-- Substition for iff
|
||||
-- ------------------
|
||||
|
||||
theorem subst_iff {P : Prop → Prop} {C : congr iff iff P} {a b : Prop} (H : a ↔ b) (H1 : P a) :
|
||||
theorem subst_iff ⦃P : Prop → Prop⦄ {C : congr iff iff P} {a b : Prop} (H : a ↔ b) (H1 : P a) :
|
||||
P b :=
|
||||
@general_operations.subst Prop iff P C a b H H1
|
||||
|
||||
-- Support for calc
|
||||
|
||||
-- Support for calculations with iff
|
||||
-- ----------------
|
||||
|
||||
calc_refl iff_refl
|
||||
calc_subst subst_iff
|
||||
calc_trans iff_trans
|
||||
calc_subst subst_iff
|
||||
|
||||
namespace iff_ops
|
||||
postfix `⁻¹`:100 := iff_symm
|
||||
infixr `⬝`:75 := iff_trans
|
||||
infixr `▸`:75 := subst_iff
|
||||
abbreviation refl := iff_refl
|
||||
abbreviation symm := iff_symm
|
||||
abbreviation trans := iff_trans
|
||||
abbreviation subst := subst_iff
|
||||
abbreviation mp := iff_mp
|
||||
end iff_ops
|
||||
|
||||
|
||||
|
||||
|
||||
|
||||
-- Boolean calculations
|
||||
-- --------------------
|
||||
|
||||
|
@ -144,4 +148,7 @@ calc
|
|||
|
||||
-- TODO: add or_left_comm, and_right_comm, and_left_comm
|
||||
|
||||
-- TODO: this is only temporary, until the calc bug is fixed
|
||||
calc_subst subst
|
||||
|
||||
end relation
|
||||
|
|
|
@ -187,21 +187,10 @@ end mp_like
|
|||
-- Notation for operations on general symbols
|
||||
-- ------------------------------------------
|
||||
|
||||
namespace general_operations
|
||||
|
||||
-- e.g. if R is an instance of the class, then "refl R" is reflexivity for the class
|
||||
definition refl := is_reflexive.infer
|
||||
definition symm := is_symmetric.infer
|
||||
definition trans := is_transitive.infer
|
||||
definition mp := mp_like.infer
|
||||
|
||||
end general_operations
|
||||
|
||||
-- namespace
|
||||
--
|
||||
-- postfix `⁻¹`:100 := operations.symm
|
||||
-- infixr `⬝`:75 := operations.trans
|
||||
|
||||
-- end symbols
|
||||
definition rel_refl := is_reflexive.infer
|
||||
definition rel_symm := is_symmetric.infer
|
||||
definition rel_trans := is_transitive.infer
|
||||
definition rel_mp := mp_like.infer
|
||||
|
||||
end relation
|
||||
|
|
|
@ -14,9 +14,9 @@ namespace lean {
|
|||
static name g_string_macro("string_macro");
|
||||
static std::string g_string_opcode("Str");
|
||||
|
||||
static expr g_bool(Const(name("bool", "bool")));
|
||||
static expr g_ff(Const(name("bool", "ff")));
|
||||
static expr g_tt(Const(name("bool", "tt")));
|
||||
static expr g_bool(Const(name("bool")));
|
||||
static expr g_ff(Const(name("ff")));
|
||||
static expr g_tt(Const(name("tt")));
|
||||
static expr g_char(Const(name("string", "char")));
|
||||
static expr g_ascii(Const(name("string", "ascii")));
|
||||
static expr g_string(Const(name("string", "string")));
|
||||
|
|
|
@ -1,4 +1,5 @@
|
|||
import logic logic.axioms.hilbert
|
||||
using inhabited nonempty
|
||||
|
||||
definition v1 : Prop := epsilon (λ x, true)
|
||||
inductive Empty : Type
|
||||
|
|
|
@ -1,5 +1,5 @@
|
|||
import standard
|
||||
using bool eq_proofs tactic
|
||||
using bool eq_ops tactic
|
||||
|
||||
variables a b c : bool
|
||||
axiom H1 : a = b
|
||||
|
|
|
@ -1,8 +1,8 @@
|
|||
import standard
|
||||
using num prod
|
||||
using num prod inhabited
|
||||
|
||||
definition H : inhabited (Prop × num × (num → num)) := _
|
||||
|
||||
(*
|
||||
print(get_env():find("H"):value())
|
||||
*)
|
||||
*)
|
||||
|
|
|
@ -1,9 +1,9 @@
|
|||
import standard
|
||||
using num prod
|
||||
using num prod nonempty inhabited
|
||||
|
||||
theorem H {A B : Type} (H1 : inhabited A) : inhabited (Prop × A × (B → num))
|
||||
:= _
|
||||
|
||||
(*
|
||||
print(get_env():find("H"):value())
|
||||
*)
|
||||
*)
|
||||
|
|
|
@ -1,5 +1,5 @@
|
|||
import standard
|
||||
using num prod
|
||||
using num prod inhabited
|
||||
|
||||
section
|
||||
parameter {A : Type}
|
||||
|
@ -16,4 +16,4 @@ end
|
|||
|
||||
(*
|
||||
print(get_env():find("tst"):value())
|
||||
*)
|
||||
*)
|
||||
|
|
|
@ -8,10 +8,10 @@ inductive t2 : Type :=
|
|||
| mk2 : t2
|
||||
|
||||
theorem inhabited_t1 : inhabited t1
|
||||
:= inhabited_intro mk1
|
||||
:= inhabited_mk mk1
|
||||
|
||||
theorem inhabited_t2 : inhabited t2
|
||||
:= inhabited_intro mk2
|
||||
:= inhabited_mk mk2
|
||||
|
||||
instance inhabited_t1 inhabited_t2
|
||||
|
||||
|
|
|
@ -44,4 +44,4 @@ function test_unify(env, m, lhs, rhs, num_s)
|
|||
end
|
||||
print("=====================")
|
||||
test_unify(get_env(), m, isNil(Nil), isNil(m), 2)
|
||||
*)
|
||||
*)
|
||||
|
|
|
@ -10,8 +10,7 @@
|
|||
-- Basic properties of lists.
|
||||
|
||||
import data.nat
|
||||
using nat eq_proofs
|
||||
|
||||
using nat eq_ops
|
||||
inductive list (T : Type) : Type :=
|
||||
| nil {} : list T
|
||||
| cons : T → list T → list T
|
||||
|
|
|
@ -16,10 +16,10 @@ theorem zero_or_succ (n : nat) : n = zero ∨ n = succ (pred n)
|
|||
:= induction_on n
|
||||
(or_intro_left _ (refl zero))
|
||||
(take m IH, or_intro_right _
|
||||
(show succ m = succ (pred (succ m)), from congr2 succ (symm (pred_succ m))))
|
||||
(show succ m = succ (pred (succ m)), from congr_arg succ (symm (pred_succ m))))
|
||||
|
||||
theorem zero_or_succ2 (n : nat) : n = zero ∨ n = succ (pred n)
|
||||
:= @induction_on _ n
|
||||
(or_intro_left _ (refl zero))
|
||||
(take m IH, or_intro_right _
|
||||
(show succ m = succ (pred (succ m)), from congr2 succ (symm (pred_succ m))))
|
||||
(show succ m = succ (pred (succ m)), from congr_arg succ (symm (pred_succ m))))
|
||||
|
|
|
@ -1,5 +1,5 @@
|
|||
import standard
|
||||
using num eq_proofs
|
||||
using num eq_ops
|
||||
|
||||
inductive nat : Type :=
|
||||
| zero : nat
|
||||
|
|
|
@ -1,5 +1,5 @@
|
|||
import standard
|
||||
using num eq_proofs
|
||||
using num eq_ops
|
||||
|
||||
inductive nat : Type :=
|
||||
| zero : nat
|
||||
|
|
|
@ -1,5 +1,5 @@
|
|||
import standard
|
||||
using num eq_proofs
|
||||
using num eq_ops
|
||||
|
||||
inductive nat : Type :=
|
||||
| zero : nat
|
||||
|
|
|
@ -1,5 +1,5 @@
|
|||
import standard
|
||||
using num eq_proofs
|
||||
using num eq_ops
|
||||
|
||||
inductive nat : Type :=
|
||||
| zero : nat
|
||||
|
|
|
@ -1,5 +1,5 @@
|
|||
import standard
|
||||
using eq_proofs
|
||||
using eq_ops
|
||||
|
||||
inductive nat : Type :=
|
||||
| zero : nat
|
||||
|
@ -16,4 +16,4 @@ variable P : nat → Prop
|
|||
|
||||
print "==========================="
|
||||
theorem tst (n : nat) (H : P (n * zero)) : P zero
|
||||
:= subst (mul_zero_right _) H
|
||||
:= subst (mul_zero_right _) H
|
||||
|
|
|
@ -4,15 +4,15 @@ using relation
|
|||
namespace is_equivalence
|
||||
|
||||
inductive class {T : Type} (R : T → T → Type) : Prop :=
|
||||
| mk : is_reflexive.class R → is_symmetric.class R → is_transitive.class R → class R
|
||||
| mk : is_reflexive R → is_symmetric R → is_transitive R → class R
|
||||
|
||||
theorem is_reflexive {T : Type} {R : T → T → Type} {C : class R} : is_reflexive.class R :=
|
||||
theorem is_reflexive {T : Type} {R : T → T → Type} {C : class R} : is_reflexive R :=
|
||||
class_rec (λx y z, x) C
|
||||
|
||||
theorem is_symmetric {T : Type} {R : T → T → Type} {C : class R} : is_symmetric.class R :=
|
||||
theorem is_symmetric {T : Type} {R : T → T → Type} {C : class R} : is_symmetric R :=
|
||||
class_rec (λx y z, y) C
|
||||
|
||||
theorem is_transitive {T : Type} {R : T → T → Type} {C : class R} : is_transitive.class R :=
|
||||
theorem is_transitive {T : Type} {R : T → T → Type} {C : class R} : is_transitive R :=
|
||||
class_rec (λx y z, z) C
|
||||
|
||||
end is_equivalence
|
||||
|
@ -25,14 +25,13 @@ theorem and_inhabited_left {a : Prop} (b : Prop) (Ha : a) : a ∧ b ↔ b :=
|
|||
iff_intro (take Hab, and_elim_right Hab) (take Hb, and_intro Ha Hb)
|
||||
|
||||
theorem test (a b c : Prop) (P : Prop → Prop) (H1 : a ↔ b) (H2 : c ∧ a) : c ∧ b :=
|
||||
gensubst.subst H1 H2
|
||||
subst_iff H1 H2
|
||||
|
||||
theorem test2 (Q R S : Prop) (H3 : R ↔ Q) (H1 : S) : Q ↔ (S ∧ Q) :=
|
||||
relation.operations.symm (and_inhabited_left Q H1)
|
||||
iff_symm (and_inhabited_left Q H1)
|
||||
|
||||
theorem test3 (Q R S : Prop) (H3 : R ↔ Q) (H1 : S) : R ↔ (S ∧ Q) :=
|
||||
gensubst.subst (test2 Q R S H3 H1) H3
|
||||
subst_iff (test2 Q R S H3 H1) H3
|
||||
|
||||
-- the composition of test2' and test3' fails
|
||||
theorem test4 (Q R S : Prop) (H3 : R ↔ Q) (H1 : S) : R ↔ (S ∧ Q) :=
|
||||
gensubst.subst (relation.operations.symm (and_inhabited_left Q H1)) H3
|
||||
subst_iff (iff_symm (and_inhabited_left Q H1)) H3
|
||||
|
|
|
@ -1,4 +1,4 @@
|
|||
import standard
|
||||
import standard data.string
|
||||
using tactic
|
||||
|
||||
variable A : Type.{1}
|
||||
|
|
|
@ -6,7 +6,7 @@ variable f : A → A → A
|
|||
|
||||
theorem tst {a b c : A} (H1 : a = b) (H2 : b = c) : f a b = f b c
|
||||
:= by apply (@congr A A (f a) (f b));
|
||||
apply (congr2 f);
|
||||
apply (congr_arg f);
|
||||
!assumption
|
||||
|
||||
|
||||
|
|
|
@ -1,19 +1,19 @@
|
|||
import standard
|
||||
using tactic
|
||||
using tactic inhabited
|
||||
|
||||
inductive sum (A : Type) (B : Type) : Type :=
|
||||
| inl : A → sum A B
|
||||
| inr : B → sum A B
|
||||
|
||||
theorem inl_inhabited {A : Type} (B : Type) (H : inhabited A) : inhabited (sum A B)
|
||||
:= inhabited_elim H (λ a, inhabited_intro (inl B a))
|
||||
:= inhabited_destruct H (λ a, inhabited_mk (inl B a))
|
||||
|
||||
theorem inr_inhabited (A : Type) {B : Type} (H : inhabited B) : inhabited (sum A B)
|
||||
:= inhabited_elim H (λ b, inhabited_intro (inr A b))
|
||||
:= inhabited_destruct H (λ b, inhabited_mk (inr A b))
|
||||
|
||||
definition my_tac := fixpoint (λ t, [ apply @inl_inhabited; t
|
||||
| apply @inr_inhabited; t
|
||||
| apply @num.inhabited_num
|
||||
| apply @num.num_inhabited
|
||||
])
|
||||
|
||||
tactic_hint [inhabited] my_tac
|
||||
|
|
|
@ -1,22 +1,22 @@
|
|||
import standard
|
||||
using tactic
|
||||
using tactic inhabited
|
||||
|
||||
inductive sum (A : Type) (B : Type) : Type :=
|
||||
| inl : A → sum A B
|
||||
| inr : B → sum A B
|
||||
|
||||
theorem inl_inhabited {A : Type} (B : Type) (H : inhabited A) : inhabited (sum A B)
|
||||
:= inhabited_elim H (λ a, inhabited_intro (inl B a))
|
||||
:= inhabited_destruct H (λ a, inhabited_mk (inl B a))
|
||||
|
||||
theorem inr_inhabited (A : Type) {B : Type} (H : inhabited B) : inhabited (sum A B)
|
||||
:= inhabited_elim H (λ b, inhabited_intro (inr A b))
|
||||
:= inhabited_destruct H (λ b, inhabited_mk (inr A b))
|
||||
|
||||
infixl `..`:100 := append
|
||||
|
||||
definition my_tac := repeat (trace "iteration"; state;
|
||||
( apply @inl_inhabited; trace "used inl"
|
||||
.. apply @inr_inhabited; trace "used inr"
|
||||
.. apply @num.inhabited_num; trace "used num")) ; now
|
||||
.. apply @num.num_inhabited; trace "used num")) ; now
|
||||
|
||||
|
||||
tactic_hint [inhabited] my_tac
|
||||
|
|
|
@ -1,5 +1,5 @@
|
|||
import standard
|
||||
using bool eq_proofs tactic
|
||||
using bool eq_ops tactic
|
||||
|
||||
variables a b c : bool
|
||||
axiom H1 : a = b
|
||||
|
|
|
@ -14,7 +14,7 @@ import logic data.nat
|
|||
|
||||
using nat
|
||||
-- using congr
|
||||
using eq_proofs
|
||||
using eq_ops
|
||||
|
||||
namespace list
|
||||
|
||||
|
|
|
@ -4,7 +4,7 @@
|
|||
-- Author: Floris van Doorn
|
||||
----------------------------------------------------------------------------------------------------
|
||||
import standard struc.binary
|
||||
using tactic num binary eq_proofs
|
||||
using tactic num binary eq_ops
|
||||
using decidable
|
||||
|
||||
namespace nat
|
||||
|
@ -57,7 +57,7 @@ theorem zero_or_succ (n : ℕ) : n = 0 ∨ n = succ (pred n)
|
|||
:= induction_on n
|
||||
(or_intro_left _ (refl 0))
|
||||
(take m IH, or_intro_right _
|
||||
(show succ m = succ (pred (succ m)), from congr2 succ (pred_succ m⁻¹)))
|
||||
(show succ m = succ (pred (succ m)), from congr_arg succ (pred_succ m⁻¹)))
|
||||
|
||||
theorem zero_or_succ2 (n : ℕ) : n = 0 ∨ ∃k, n = succ k
|
||||
:= or_imp_or (zero_or_succ n) (assume H, H) (assume H : n = succ (pred n), exists_intro (pred n) H)
|
||||
|
@ -96,7 +96,7 @@ theorem decidable_eq [instance] (n m : ℕ) : decidable (n = m)
|
|||
(λ (n' : ℕ) (iH2 : decidable (n' = succ m')),
|
||||
have d1 : decidable (n' = m'), from iH1 n',
|
||||
decidable.rec_on d1
|
||||
(assume Heq : n' = m', inl (congr2 succ Heq))
|
||||
(assume Heq : n' = m', inl (congr_arg succ Heq))
|
||||
(assume Hne : n' ≠ m',
|
||||
have H1 : succ n' ≠ succ m', from
|
||||
assume Heq, absurd (succ_inj Heq) Hne,
|
||||
|
|
Loading…
Reference in a new issue