refactor(logic/funext.lean, algebra/function.lean): delete logic/funext, merge into algebra/function
This commit is contained in:
parent
f65a49b2c3
commit
db7bdce451
9 changed files with 27 additions and 41 deletions
|
@ -6,7 +6,7 @@ Module: algebra.category.functor
|
||||||
Author: Floris van Doorn
|
Author: Floris van Doorn
|
||||||
-/
|
-/
|
||||||
import .basic
|
import .basic
|
||||||
import logic.cast logic.funext
|
import logic.cast algebra.function
|
||||||
open function
|
open function
|
||||||
open category eq eq.ops heq
|
open category eq eq.ops heq
|
||||||
|
|
||||||
|
|
|
@ -7,6 +7,8 @@ Author: Leonardo de Moura
|
||||||
|
|
||||||
General operations on functions.
|
General operations on functions.
|
||||||
-/
|
-/
|
||||||
|
import logic.cast
|
||||||
|
|
||||||
namespace function
|
namespace function
|
||||||
|
|
||||||
variables {A : Type} {B : Type} {C : Type} {D : Type} {E : Type}
|
variables {A : Type} {B : Type} {C : Type} {D : Type} {E : Type}
|
||||||
|
@ -95,6 +97,23 @@ have h : f a = b, from calc
|
||||||
... = id b : by rewrite (right_inv_eq inv)
|
... = id b : by rewrite (right_inv_eq inv)
|
||||||
... = b : rfl,
|
... = b : rfl,
|
||||||
exists.intro a h
|
exists.intro a h
|
||||||
|
|
||||||
|
theorem compose.assoc (f : C → D) (g : B → C) (h : A → B) : (f ∘ g) ∘ h = f ∘ (g ∘ h) :=
|
||||||
|
funext (take x, rfl)
|
||||||
|
|
||||||
|
theorem compose.left_id (f : A → B) : id ∘ f = f :=
|
||||||
|
funext (take x, rfl)
|
||||||
|
|
||||||
|
theorem compose.right_id (f : A → B) : f ∘ id = f :=
|
||||||
|
funext (take x, rfl)
|
||||||
|
|
||||||
|
theorem compose_const_right (f : B → C) (b : B) : f ∘ (const A b) = const A (f b) :=
|
||||||
|
funext (take x, rfl)
|
||||||
|
|
||||||
|
theorem hfunext {A : Type} {B : A → Type} {B' : A → Type} {f : Π x, B x} {g : Π x, B' x}
|
||||||
|
(H : ∀ a, f a == g a) : f == g :=
|
||||||
|
let HH : B = B' := (funext (λ x, heq.type_eq (H x))) in
|
||||||
|
cast_to_heq (funext (λ a, heq.to_eq (heq.trans (cast_app HH f a) (H a))))
|
||||||
end function
|
end function
|
||||||
|
|
||||||
-- copy reducible annotations to top-level
|
-- copy reducible annotations to top-level
|
||||||
|
|
|
@ -1,31 +0,0 @@
|
||||||
/-
|
|
||||||
Copyright (c) 2015 Microsoft Corporation. All rights reserved.
|
|
||||||
Released under Apache 2.0 license as described in the file LICENSE.
|
|
||||||
|
|
||||||
Author: Leonardo de Moura
|
|
||||||
|
|
||||||
Basic theorems for functions
|
|
||||||
-/
|
|
||||||
import logic.cast algebra.function data.sigma
|
|
||||||
open function eq.ops
|
|
||||||
|
|
||||||
namespace function
|
|
||||||
variables {A B C D: Type}
|
|
||||||
|
|
||||||
theorem compose.assoc (f : C → D) (g : B → C) (h : A → B) : (f ∘ g) ∘ h = f ∘ (g ∘ h) :=
|
|
||||||
funext (take x, rfl)
|
|
||||||
|
|
||||||
theorem compose.left_id (f : A → B) : id ∘ f = f :=
|
|
||||||
funext (take x, rfl)
|
|
||||||
|
|
||||||
theorem compose.right_id (f : A → B) : f ∘ id = f :=
|
|
||||||
funext (take x, rfl)
|
|
||||||
|
|
||||||
theorem compose_const_right (f : B → C) (b : B) : f ∘ (const A b) = const A (f b) :=
|
|
||||||
funext (take x, rfl)
|
|
||||||
|
|
||||||
theorem hfunext {A : Type} {B : A → Type} {B' : A → Type} {f : Π x, B x} {g : Π x, B' x}
|
|
||||||
(H : ∀ a, f a == g a) : f == g :=
|
|
||||||
let HH : B = B' := (funext (λ x, heq.type_eq (H x))) in
|
|
||||||
cast_to_heq (funext (λ a, heq.to_eq (heq.trans (cast_app HH f a) (H a))))
|
|
||||||
end function
|
|
|
@ -1,5 +1,4 @@
|
||||||
import algebra.function
|
import algebra.function
|
||||||
import logic.funext
|
|
||||||
|
|
||||||
open function
|
open function
|
||||||
|
|
||||||
|
|
|
@ -1,4 +1,4 @@
|
||||||
550.lean:44:72: error:invalid 'rewrite' tactic, step produced type incorrect term, details: type mismatch at application
|
550.lean:43:72: error:invalid 'rewrite' tactic, step produced type incorrect term, details: type mismatch at application
|
||||||
eq.symm linv
|
eq.symm linv
|
||||||
term
|
term
|
||||||
linv
|
linv
|
||||||
|
@ -27,14 +27,14 @@ rinv : func ∘ finv = function.id
|
||||||
(eq.symm rinv))
|
(eq.symm rinv))
|
||||||
(eq.symm (function.compose.assoc finv func finv)))
|
(eq.symm (function.compose.assoc finv func finv)))
|
||||||
(function.compose.assoc (finv ∘ func) finv func)) = id
|
(function.compose.assoc (finv ∘ func) finv func)) = id
|
||||||
550.lean:44:47: error: don't know how to synthesize placeholder
|
550.lean:43:47: error: don't know how to synthesize placeholder
|
||||||
A : Type,
|
A : Type,
|
||||||
f : bijection A,
|
f : bijection A,
|
||||||
func finv : A → A,
|
func finv : A → A,
|
||||||
linv : finv ∘ func = function.id,
|
linv : finv ∘ func = function.id,
|
||||||
rinv : func ∘ finv = function.id
|
rinv : func ∘ finv = function.id
|
||||||
⊢ inv (mk func finv linv rinv) ∘b mk func finv linv rinv = id
|
⊢ inv (mk func finv linv rinv) ∘b mk func finv linv rinv = id
|
||||||
550.lean:44:2: error: failed to add declaration 'bijection.inv.linv' to environment, value has metavariables
|
550.lean:43:2: error: failed to add declaration 'bijection.inv.linv' to environment, value has metavariables
|
||||||
remark: set 'formatter.hide_full_terms' to false to see the complete term
|
remark: set 'formatter.hide_full_terms' to false to see the complete term
|
||||||
λ (A : Type) (f : …),
|
λ (A : Type) (f : …),
|
||||||
bijection.rec_on f (λ (func finv : …) (linv : …) (rinv : …), ?M_1)
|
bijection.rec_on f (λ (func finv : …) (linv : …) (rinv : …), ?M_1)
|
||||||
|
|
|
@ -6,5 +6,5 @@ check ∅ -- o.k.
|
||||||
check λs t, subset s t -- o.k.
|
check λs t, subset s t -- o.k.
|
||||||
check λs t, s ⊆ t -- fixed
|
check λs t, s ⊆ t -- fixed
|
||||||
|
|
||||||
infix `⊆`:50 := subset
|
infix `⊆` := subset
|
||||||
check λs t, s ⊆ t
|
check λs t, s ⊆ t
|
||||||
|
|
|
@ -12,8 +12,8 @@ definition set.subset {A : Type₁} (s₁ s₂ : set A) : Prop :=
|
||||||
definition finset.subset {A : Type₁} (s₁ s₂ : finset A) : Prop :=
|
definition finset.subset {A : Type₁} (s₁ s₂ : finset A) : Prop :=
|
||||||
∀ ⦃a : A⦄, a ∈ s₁ → a ∈ s₂
|
∀ ⦃a : A⦄, a ∈ s₁ → a ∈ s₂
|
||||||
|
|
||||||
infix `⊆`:50 := set.subset
|
infix `⊆` := set.subset
|
||||||
infix `⊆`:50 := finset.subset
|
infix `⊆` := finset.subset
|
||||||
|
|
||||||
example (A : Type₁) (x : A) (S H : set A) (Pin : x ∈ S) (Psub : S ⊆ H) : x ∈ H :=
|
example (A : Type₁) (x : A) (S H : set A) (Pin : x ∈ S) (Psub : S ⊆ H) : x ∈ H :=
|
||||||
Psub Pin -- Error, we cannot infer at preprocessing time the binder information for Psub
|
Psub Pin -- Error, we cannot infer at preprocessing time the binder information for Psub
|
||||||
|
|
|
@ -1,5 +1,4 @@
|
||||||
import algebra.function
|
import algebra.function
|
||||||
import logic.funext
|
|
||||||
|
|
||||||
open function
|
open function
|
||||||
|
|
||||||
|
|
|
@ -1,4 +1,4 @@
|
||||||
tactic_id_bug.lean:23:4: proof state
|
tactic_id_bug.lean:22:4: proof state
|
||||||
A : Type,
|
A : Type,
|
||||||
gfunc gfinv : A → A,
|
gfunc gfinv : A → A,
|
||||||
glinv : gfinv ∘ gfunc = id,
|
glinv : gfinv ∘ gfunc = id,
|
||||||
|
|
Loading…
Reference in a new issue