Skip to content

Commit b0c3952

Browse files
committed
refactor(Mathport): move mathport prelude to mathlib4 (#80)
Moving `Mathport/Prelude/*` out of mathport into mathlib4. I've added a README that will hopefully help bystanders: ``` Mathport prelude === This directory contains instructions for `mathport`, helping it to translate `mathlib` and align declarations and tactics with `mathlib4`. (These files were formerly part of `mathport`, in the directory `Mathport/Prelude/`.) `SpecialNames.lean` : Contains `#align X Y` statements, where `X` is an identifier from mathlib3 which should be aligned with the identifier `Y` from mathlib4. Sometimes we need `#align` statements just to handle exceptions to casing rules, but there are also many exceptional cases. `Syntax.lean` : Contains unimplemented stubs of tactics which need to be migrated from Lean3 to Lean4. When porting tactics, you can move the relevant stubs to a new file and use them as a starting point. Please make sure this file stays in sync with new tactic implementations (and in particular that the syntax is not defined twice). Please preserve the syntax of existing mathlib tactics, so that there are no unnecessary parse errors in the source files generated by `synport`. It is fine to fail with an error message for unimplemented features for now. ``` Co-authored-by: Scott Morrison <scott.morrison@gmail.com>
1 parent 28111cf commit b0c3952

18 files changed

+1002
-26
lines changed

.github/workflows/build.yml

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -25,7 +25,7 @@ jobs:
2525

2626
- name: update Mathlib.lean
2727
run: |
28-
find Mathlib -not -type d | sort | sed 's/\.lean//;s,/,.,g;s/^/import /' > Mathlib.lean
28+
find Mathlib -name "*.lean" | sort | sed 's/\.lean//;s,/,.,g;s/^/import /' > Mathlib.lean
2929
3030
- name: build mathlib
3131
run: lake build

Mathlib.lean

Lines changed: 5 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -30,6 +30,10 @@ import Mathlib.Init.SetNotation
3030
import Mathlib.Lean.LocalContext
3131
import Mathlib.Logic.Basic
3232
import Mathlib.Logic.Function.Basic
33+
import Mathlib.Mathport.Attributes
34+
import Mathlib.Mathport.Rename
35+
import Mathlib.Mathport.SpecialNames
36+
import Mathlib.Mathport.Syntax
3337
import Mathlib.Tactic.Basic
3438
import Mathlib.Tactic.Cache
3539
import Mathlib.Tactic.Coe
@@ -41,6 +45,7 @@ import Mathlib.Tactic.NoMatch
4145
import Mathlib.Tactic.NormNum
4246
import Mathlib.Tactic.OpenPrivate
4347
import Mathlib.Tactic.PrintPrefix
48+
import Mathlib.Tactic.Rcases
4449
import Mathlib.Tactic.Ring
4550
import Mathlib.Tactic.RunTac
4651
import Mathlib.Tactic.ShowTerm

Mathlib/Mathport/Attributes.lean

Lines changed: 19 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,19 @@
1+
/-
2+
Copyright (c) 2021 Microsoft Corporation. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Mario Carneiro
5+
-/
6+
import Lean.Attributes
7+
8+
namespace Lean.Attr
9+
10+
initialize reflAttr : TagAttribute ← registerTagAttribute `refl "reflexive relation"
11+
initialize symmAttr : TagAttribute ← registerTagAttribute `symm "symmetric relation"
12+
initialize transAttr : TagAttribute ← registerTagAttribute `trans "transitive relation"
13+
initialize substAttr : TagAttribute ← registerTagAttribute `subst "substitution"
14+
15+
initialize linterAttr : TagAttribute ←
16+
registerTagAttribute `linter "Use this declaration as a linting test in #lint"
17+
18+
initialize hintTacticAttr : TagAttribute ←
19+
registerTagAttribute `hintTactic "A tactic that should be tried by `hint`."

Mathlib/Mathport/README.md

Lines changed: 22 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,22 @@
1+
Mathport prelude
2+
===
3+
4+
This directory contains instructions for `mathport`,
5+
helping it to translate `mathlib` and align declarations and tactics with `mathlib4`.
6+
(These files were formerly part of `mathport`, in the directory `Mathport/Prelude/`.)
7+
8+
`SpecialNames.lean`
9+
: Contains `#align X Y` statements, where `X` is an identifier from mathlib3
10+
which should be aligned with the identifier `Y` from mathlib4.
11+
Sometimes we need `#align` statements just to handle exceptions to casing rules,
12+
but there are also many exceptional cases.
13+
14+
`Syntax.lean`
15+
: Contains unimplemented stubs of tactics which need to be migrated from Lean3 to Lean4.
16+
When porting tactics, you can move the relevant stubs to a new file and
17+
use them as a starting point.
18+
Please make sure this file stays in sync with new tactic implementations
19+
(and in particular that the syntax is not defined twice).
20+
Please preserve the syntax of existing mathlib tactics,
21+
so that there are no unnecessary parse errors in the source files generated by `synport`.
22+
It is fine to fail with an error message for unimplemented features for now.

Mathlib/Mathport/Rename.lean

Lines changed: 51 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,51 @@
1+
/-
2+
Copyright (c) 2021 Microsoft Corporation. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Daniel Selsam
5+
-/
6+
import Lean
7+
8+
namespace Mathlib.Prelude.Rename
9+
10+
open Lean
11+
open System (FilePath)
12+
open Std (HashMap)
13+
14+
abbrev RenameMap := HashMap Name Name
15+
16+
def RenameMap.insertPair (m : RenameMap) : Name × Name → RenameMap
17+
| (n3, n4) => m.insert n3 n4
18+
19+
initialize renameExtension : SimplePersistentEnvExtension (Name × Name) RenameMap ←
20+
registerSimplePersistentEnvExtension {
21+
name := `renameMapExtension
22+
addEntryFn := RenameMap.insertPair
23+
addImportedFn := fun es => mkStateFromImportedEntries (RenameMap.insertPair) {} es
24+
}
25+
26+
def getRenameMap (env : Environment) : RenameMap := do
27+
renameExtension.getState env
28+
29+
def addNameAlignment (n3 n4 : Name) : CoreM Unit := do
30+
modifyEnv fun env => renameExtension.addEntry env (n3, n4)
31+
32+
open Lean.Elab Lean.Elab.Command
33+
34+
syntax (name := align) "#align " ident ident : command
35+
36+
@[commandElab align] def elabAlign : CommandElab
37+
| `(#align%$tk $id3:ident $id4:ident) =>
38+
liftCoreM $ addNameAlignment id3.getId id4.getId
39+
| _ => throwUnsupportedSyntax
40+
41+
syntax (name := lookup3) "#lookup3 " ident : command
42+
43+
@[commandElab lookup3] def elabLookup3 : CommandElab
44+
| `(#lookup3%$tk $id3:ident) => do
45+
let n3 := id3.getId
46+
match getRenameMap (← getEnv) |>.find? n3 with
47+
| none => logInfoAt tk s!"name `{n3} not found"
48+
| some n4 => logInfoAt tk s!"{n4}"
49+
| _ => throwUnsupportedSyntax
50+
51+
end Mathlib.Prelude.Rename

Mathlib/Mathport/SpecialNames.lean

Lines changed: 73 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,73 @@
1+
/-
2+
Copyright (c) 2021 Microsoft Corporation. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Daniel Selsam
5+
-/
6+
import Mathlib.Mathport.Rename
7+
8+
namespace Mathlib.Prelude
9+
10+
-- Note: we do not currently auto-align constants.
11+
#align quot Quot
12+
#align quot.mk Quot.mk
13+
#align quot.lift Quot.lift
14+
#align quot.ind Quot.ind
15+
16+
-- Otherwise would be`OutParam` and `OptParam`!
17+
-- Note: we want `auto_param` to *not* align.
18+
#align out_param outParam
19+
#align opt_param optParam
20+
21+
#align heq HEq
22+
23+
#align punit PUnit
24+
#align pprod PProd
25+
26+
#align has_coe Coe
27+
#align has_coe.coe Coe.coe
28+
29+
#align has_coe_t CoeT
30+
#align has_coe_t.coe CoeT.coe
31+
32+
#align has_coe_to_fun CoeFun
33+
#align has_coe_to_fun.coe CoeFun.coe
34+
35+
#align has_coe_to_sort CoeSort
36+
#align has_coe_to_sort.coe CoeSort.coe
37+
38+
#align has_le LE
39+
#align has_le.le LE.le
40+
#align has_lt LT
41+
#align has_lt.lt LT.lt
42+
43+
#align has_beq BEq
44+
#align has_sizeof SizeOf
45+
46+
-- This otherwise causes filenames to clash
47+
#align category_theory.Kleisli CategoryTheory.KleisliCat
48+
49+
-- TODO: backport?
50+
#align int.neg_succ_of_nat Int.negSucc
51+
52+
-- Generic 'has'-stripping
53+
-- Note: we don't currently strip automatically for various reasons.
54+
#align has_add Add
55+
#align has_sub Sub
56+
#align has_mul Mul
57+
#align has_div Div
58+
#align has_neg Neg
59+
#align has_mod Mod
60+
#align has_pow Pow
61+
#align has_append Append
62+
#align has_pure Pure
63+
#align has_bind Bind
64+
#align has_seq Seq
65+
#align has_seq_left SeqLeft
66+
#align has_seq_right SeqRight
67+
68+
-- Implementation detail
69+
#align _sorry_placeholder_ _sorry_placeholder_
70+
71+
72+
73+
end Mathlib.Prelude

0 commit comments

Comments
 (0)