├── lean-toolchain ├── nix ├── .gitattributes ├── flake.nix └── flake.lock ├── LottExamples ├── Lambda │ ├── lambda.lottinput │ ├── eta.lottinput │ ├── intro.tex.lotttmpl │ ├── Tex.lean │ ├── Syntax.lean │ └── Semantics.lean ├── STLC │ ├── Tex.lean │ ├── Syntax.lean │ ├── Semantics.lean │ └── Proofs.lean ├── SystemF │ ├── Tex.lean │ ├── Syntax.lean │ └── Semantics.lean ├── Tex.lean ├── STLC.lean ├── Lambda.lean ├── SystemF.lean ├── .gitignore ├── Makefile └── main.tex ├── .envrc ├── .gitignore ├── Lott ├── Environment.lean ├── Elab.lean ├── Data │ ├── Function.lean │ ├── Char.lean │ ├── Substring.lean │ ├── Nat.lean │ ├── Name.lean │ ├── Option.lean │ ├── String.lean │ ├── List.lean │ └── Range.lean ├── Parser │ ├── NotJudgement.lean │ ├── OrJudgement.lean │ ├── ImpliesJudgement.lean │ ├── NotExistentialJudgement.lean │ ├── JudgementComprehension.lean │ ├── UniversalJudgement.lean │ └── Filter.lean ├── Environment │ ├── MetaVar.lean │ └── Basic.lean ├── Elab │ ├── NotJudgement.lean │ ├── Bool.lean │ ├── OrJudgement.lean │ ├── ImpliesJudgement.lean │ ├── NotExistentialJudgement.lean │ ├── Options.lean │ ├── JudgementComprehension.lean │ ├── Nat.lean │ ├── UniversalJudgement.lean │ ├── Filter.lean │ ├── MetaVar.lean │ ├── Judgement.lean │ ├── Basic.lean │ └── NonTerminal.lean ├── Tactic │ └── Termination.lean ├── Linter │ └── UnusedVariables.lean ├── Parser.lean └── IR.lean ├── LottExamples.lean ├── Lott.lean ├── .github └── workflows │ └── build.yaml ├── Makefile ├── README.md ├── lake-manifest.json └── lakefile.lean /lean-toolchain: -------------------------------------------------------------------------------- 1 | leanprover/lean4:v4.17.0 2 | -------------------------------------------------------------------------------- /nix/.gitattributes: -------------------------------------------------------------------------------- 1 | /flake.lock -diff 2 | -------------------------------------------------------------------------------- /LottExamples/Lambda/lambda.lottinput: -------------------------------------------------------------------------------- 1 | λ x. e 2 | -------------------------------------------------------------------------------- /.envrc: -------------------------------------------------------------------------------- 1 | use flake "git+file://$PWD?dir=nix#tex" 2 | -------------------------------------------------------------------------------- /.gitignore: -------------------------------------------------------------------------------- 1 | /.direnv 2 | result 3 | 4 | /.lake 5 | -------------------------------------------------------------------------------- /Lott/Environment.lean: -------------------------------------------------------------------------------- 1 | import Lott.Environment.Basic 2 | -------------------------------------------------------------------------------- /LottExamples/Lambda/eta.lottinput: -------------------------------------------------------------------------------- 1 | λ x. e x$0 ->η e 2 | -------------------------------------------------------------------------------- /LottExamples/STLC/Tex.lean: -------------------------------------------------------------------------------- 1 | import LottExamples.STLC.Semantics 2 | import LottExamples.STLC.Syntax 3 | -------------------------------------------------------------------------------- /LottExamples.lean: -------------------------------------------------------------------------------- 1 | import LottExamples.Lambda 2 | import LottExamples.STLC 3 | import LottExamples.SystemF 4 | -------------------------------------------------------------------------------- /LottExamples/SystemF/Tex.lean: -------------------------------------------------------------------------------- 1 | import LottExamples.SystemF.Semantics 2 | import LottExamples.SystemF.Syntax 3 | -------------------------------------------------------------------------------- /Lott.lean: -------------------------------------------------------------------------------- 1 | import Lott.Elab 2 | import Lott.Environment 3 | import Lott.Linter.UnusedVariables 4 | import Lott.Parser 5 | -------------------------------------------------------------------------------- /LottExamples/Tex.lean: -------------------------------------------------------------------------------- 1 | import LottExamples.Lambda.Tex 2 | import LottExamples.STLC.Tex 3 | import LottExamples.SystemF.Tex 4 | -------------------------------------------------------------------------------- /LottExamples/STLC.lean: -------------------------------------------------------------------------------- 1 | import LottExamples.STLC.Proofs 2 | import LottExamples.STLC.Semantics 3 | import LottExamples.STLC.Syntax 4 | -------------------------------------------------------------------------------- /Lott/Elab.lean: -------------------------------------------------------------------------------- 1 | import Lott.Elab.Basic 2 | import Lott.Elab.Judgement 3 | import Lott.Elab.MetaVar 4 | import Lott.Elab.NonTerminal 5 | -------------------------------------------------------------------------------- /LottExamples/Lambda.lean: -------------------------------------------------------------------------------- 1 | import LottExamples.Lambda.Semantics 2 | import LottExamples.Lambda.Syntax 3 | import LottExamples.Lambda.Tex 4 | -------------------------------------------------------------------------------- /LottExamples/Lambda/intro.tex.lotttmpl: -------------------------------------------------------------------------------- 1 | This is a test of embedding symbols like [[λ x. e]] and judgements like [[λ x. e x$0 ->η e]]. 2 | -------------------------------------------------------------------------------- /Lott/Data/Function.lean: -------------------------------------------------------------------------------- 1 | namespace Function 2 | 3 | def Injective' (f : α → β) : Prop := ∀ x y, f x = f y → x = y 4 | 5 | end Function 6 | -------------------------------------------------------------------------------- /LottExamples/SystemF.lean: -------------------------------------------------------------------------------- 1 | import LottExamples.SystemF.Proofs 2 | import LottExamples.SystemF.Semantics 3 | import LottExamples.SystemF.Syntax 4 | -------------------------------------------------------------------------------- /Lott/Data/Char.lean: -------------------------------------------------------------------------------- 1 | def Char.isSubscriptAlpha (c : Char) : Bool := "ᴬᴮꟲᴰᴱꟳᴳᴴᴵᴶᴷᴸᴹᴺᴼᴾꟴᴿᵀᵁⱽᵂ𐞄𐞒𐞖ᶦᶫᶰ𐞪ᶸ𐞲ᵃᵇᶜᵈᵉᶠᵍʰⁱʲᵏˡᵐⁿᵒᵖ𐞥ʳˢᵗᵘᵛʷˣʸᶻₐₑₕᵢⱼₖₗₘₙₒₚᵣₛₜᵤᵥₓ".contains c 2 | -------------------------------------------------------------------------------- /LottExamples/Lambda/Tex.lean: -------------------------------------------------------------------------------- 1 | import Lott.Elab.Filter 2 | import LottExamples.Lambda.Semantics 3 | import LottExamples.Lambda.Syntax 4 | 5 | #filter "intro.tex.lotttmpl" 6 | -------------------------------------------------------------------------------- /Lott/Parser/NotJudgement.lean: -------------------------------------------------------------------------------- 1 | import Lott.Parser 2 | 3 | namespace Lott 4 | 5 | syntax (name := notJudgement) "¬" withPosition(Lott.Judgement) : Lott.Judgement 6 | 7 | end Lott 8 | -------------------------------------------------------------------------------- /Lott/Parser/OrJudgement.lean: -------------------------------------------------------------------------------- 1 | import Lott.Parser 2 | 3 | namespace Lott 4 | 5 | syntax (name := orJudgement) Lott.Judgement " ∨ " withPosition(Lott.Judgement) : Lott.Judgement 6 | 7 | end Lott 8 | -------------------------------------------------------------------------------- /Lott/Parser/ImpliesJudgement.lean: -------------------------------------------------------------------------------- 1 | import Lott.Parser 2 | 3 | namespace Lott 4 | 5 | syntax (name := impliesJudgement) Lott.Judgement " ⇒ " withPosition(Lott.Judgement) : Lott.Judgement 6 | 7 | end Lott 8 | -------------------------------------------------------------------------------- /LottExamples/.gitignore: -------------------------------------------------------------------------------- 1 | *.aux 2 | *.d 3 | *.fdb_latexmk 4 | *.fls 5 | *.log 6 | *.out 7 | *.pdf 8 | *.synctex 9 | *.synctex(busy) 10 | *.synctex.gz 11 | *.synctex.gz(busy) 12 | 13 | **/*.tex 14 | -------------------------------------------------------------------------------- /LottExamples/Makefile: -------------------------------------------------------------------------------- 1 | -include main.d 2 | main.pdf: main.tex 3 | latexmk -f -pdf -file-line-error -interaction=nonstopmode -deps-out=main.d $< 4 | 5 | clean: 6 | latexmk -c 7 | rm -f main.d main.pdf 8 | -------------------------------------------------------------------------------- /Lott/Data/Substring.lean: -------------------------------------------------------------------------------- 1 | partial 2 | def Substring.dropPrefixes (s pre : Substring) : Substring := 3 | if pre.bsize = 0 then 4 | s 5 | else 6 | match s.dropPrefix? pre with 7 | | some s' => s'.dropPrefixes pre 8 | | none => s 9 | -------------------------------------------------------------------------------- /Lott/Parser/NotExistentialJudgement.lean: -------------------------------------------------------------------------------- 1 | import Lott.Parser 2 | 3 | namespace Lott 4 | 5 | open Lean 6 | 7 | syntax (name := notExistentialJudgement) "∄ " explicitBinders ", " withPosition(Lott.Judgement) : Lott.Judgement 8 | 9 | end Lott 10 | -------------------------------------------------------------------------------- /.github/workflows/build.yaml: -------------------------------------------------------------------------------- 1 | name: Build 2 | on: 3 | push: 4 | branches: ["**"] 5 | jobs: 6 | build: 7 | runs-on: ubuntu-latest 8 | steps: 9 | - uses: actions/checkout@v4 10 | - uses: leanprover/lean-action@v1 11 | with: 12 | build-args: Lott LottExamples 13 | -------------------------------------------------------------------------------- /Lott/Data/Nat.lean: -------------------------------------------------------------------------------- 1 | namespace Nat 2 | 3 | theorem eq_of_beq_eq_true' {a b : Nat} (h : (a == b) = true) : a = b := by 4 | dsimp [BEq.beq, decide, instDecidableEqNat, Nat.decEq] at h 5 | split at h 6 | · case _ h' => exact Nat.eq_of_beq_eq_true h' 7 | · case _ h' => nomatch h 8 | 9 | end Nat 10 | -------------------------------------------------------------------------------- /Lott/Parser/JudgementComprehension.lean: -------------------------------------------------------------------------------- 1 | import Lott.Parser 2 | 3 | namespace Lott 4 | 5 | open Lean.Parser 6 | open Lean.Parser.Term 7 | 8 | syntax (name := judgementComprehension) "" : Lott.Judgement 9 | 10 | end Lott 11 | -------------------------------------------------------------------------------- /LottExamples/Lambda/Syntax.lean: -------------------------------------------------------------------------------- 1 | import Lott 2 | 3 | namespace LottExamples.Lambda 4 | 5 | locally_nameless 6 | metavar Var, x 7 | 8 | nonterminal Term, e := 9 | | x : var 10 | | "λ " x ". " e : lam (bind x in e) 11 | | e₀ e₁ : app 12 | | "(" e ")" : paren notex (expand := return e) 13 | 14 | end LottExamples.Lambda 15 | -------------------------------------------------------------------------------- /Makefile: -------------------------------------------------------------------------------- 1 | all: lott lott-examples 2 | 3 | .PHONY: lott 4 | lott: 5 | lake build Lott 6 | 7 | .PHONY: lott-examples 8 | lott-examples: 9 | lake build -R LottExamples LottExamples.Tex 10 | 11 | .PHONY: lott-examples-tex 12 | lott-examples-noterm: 13 | lake build -R LottExamples.Tex -K noterm=1 14 | 15 | .PHONY: lott-examples 16 | lott-examples-notex: 17 | lake build -R LottExamples -K notex=1 18 | -------------------------------------------------------------------------------- /Lott/Parser/UniversalJudgement.lean: -------------------------------------------------------------------------------- 1 | import Lott.Parser 2 | 3 | namespace Lott 4 | 5 | open Lean.Parser.Term 6 | 7 | syntax (name := universalJudgement) "∀ " (ppSpace (binderIdent <|> bracketedBinder))+ optType ", " withPosition(Lott.Judgement) : Lott.Judgement 8 | 9 | syntax (name := universalPredJudgement) "∀ " ident binderPred ", " withPosition(Lott.Judgement) : Lott.Judgement 10 | 11 | end Lott 12 | -------------------------------------------------------------------------------- /Lott/Environment/MetaVar.lean: -------------------------------------------------------------------------------- 1 | import Lean 2 | 3 | namespace Lott 4 | 5 | open Lean 6 | 7 | initialize metaVarExt : PersistentEnvExtension (Name × Bool) (Name × Bool) (NameMap Bool) ← 8 | registerPersistentEnvExtension { 9 | mkInitial := return default 10 | addImportedFn := fun metaVarss => 11 | return metaVarss.flatten.foldl (init := mkNameMap _) (fun acc (n, ln) => acc.insert n ln) 12 | addEntryFn := fun st (n, ln) => st.insert n ln 13 | exportEntriesFn := RBMap.fold (cmp := Name.quickCmp) (init := #[]) fun acc n ln => acc.push (n, ln) 14 | } 15 | 16 | end Lott 17 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # lott 2 | 3 | An [ott](https://github.com/ott-lang/ott)-like DSL embedded in Lean. 4 | 5 | ![lott](https://github.com/user-attachments/assets/671b89f7-fba0-4a90-bfed-dd01c0f8c9db) 6 | 7 | See the [`LottExamples`](LottExamples) directory for usage examples. 8 | 9 | ## References 10 | 11 | - Peter Sewell, Francesco Zappa Nardelli, Scott Owens, Gilles Peskine, Thomas Ridge, Susmit Sarkar, and others. 2010. Ott: Effective tool support for the working semanticist. Journal of functional programming 20, 1 (2010), 71–122. 12 | - Arthur Charguéraud. 2012. The locally nameless representation. Journal of automated reasoning 49, (2012), 363–408. 13 | -------------------------------------------------------------------------------- /Lott/Elab/NotJudgement.lean: -------------------------------------------------------------------------------- 1 | import Lott.Elab.Basic 2 | import Lott.Parser.NotJudgement 3 | 4 | namespace Lott 5 | 6 | open Lean 7 | open Lean.Elab 8 | 9 | @[macro Lott.judgementEmbed] 10 | private 11 | def notJudgementImpl : Macro := fun stx => do 12 | let `([[¬$j:Lott.Judgement]]) := stx | Macro.throwUnsupported 13 | ``(¬[[$j:Lott.Judgement]]) 14 | 15 | @[lott_tex_elab notJudgement] 16 | private 17 | def notJudgementTexElab : TexElab := fun profile ref stx => do 18 | let `(Lott.Judgement| ¬$j:Lott.Judgement) := stx | throwUnsupportedSyntax 19 | let jTex ← texElabSymbolOrJudgement j.raw.getKind profile ref j 20 | return s!"\\lottsym\{¬} \\, {jTex}" 21 | 22 | end Lott 23 | -------------------------------------------------------------------------------- /Lott/Elab/Bool.lean: -------------------------------------------------------------------------------- 1 | import Lott.Elab 2 | 3 | namespace Lott 4 | 5 | open Lean 6 | open Lean.Elab 7 | open Lean.Parser 8 | 9 | declare_syntax_cat Lott.Symbol.Bool 10 | 11 | run_cmd setEnv <| aliasExt.addEntry (← getEnv) { canon := `Bool, alias := `b, tex? := none } 12 | 13 | @[Lott.Symbol.Bool_parser] 14 | private 15 | def bool.b_parser : Parser := leadingNode `Lott.Symbol.Bool maxPrec <| identPrefix "b" 16 | 17 | @[macro symbolEmbed] 18 | private 19 | def boolImpl : Macro := fun stx => do 20 | let .node _ `Lott.symbolEmbed #[ 21 | .atom _ "[[", 22 | .node _ `Lott.Symbol.Bool #[b@(.ident ..)], 23 | .atom _ "]]" 24 | ] := stx | Macro.throwUnsupported 25 | return b 26 | 27 | end Lott 28 | -------------------------------------------------------------------------------- /LottExamples/Lambda/Semantics.lean: -------------------------------------------------------------------------------- 1 | import LottExamples.Lambda.Syntax 2 | 3 | namespace LottExamples.Lambda 4 | 5 | -- Alpha reduction is not included since locally nameless representation makes it meaningless. 6 | 7 | judgement_syntax e " ->ᵦ " e' : BetaReduction (tex := s!"{e} \\, \\lottsym\{\\stackrel\{β}\{→}} \\, {e'}") 8 | 9 | judgement BetaReduction where 10 | 11 | ───────────────────────── mk 12 | (λ x. e₀) e₁ ->ᵦ e₀^^e₁/x 13 | 14 | judgement_syntax "lc" "(" e ")" : Term.VarLocallyClosed 15 | 16 | judgement_syntax e " ->η " e' : EtaReduction (tex := s!"{e} \\, \\lottsym\{\\stackrel\{η}\{→}} \\, {e'}") 17 | 18 | judgement EtaReduction where 19 | 20 | lc(e) 21 | ──────────────── mk 22 | λ x. e x$0 ->η e 23 | 24 | end LottExamples.Lambda 25 | -------------------------------------------------------------------------------- /LottExamples/STLC/Syntax.lean: -------------------------------------------------------------------------------- 1 | import Lott 2 | 3 | namespace LottExamples.STLC 4 | 5 | nonterminal «Type», τ := 6 | | τ₀ " → " τ₁ : arr 7 | | "Unit" : unit 8 | 9 | locally_nameless 10 | metavar Var, x 11 | 12 | nonterminal Term, e := 13 | | x : var 14 | | "λ " x ". " e : lam (bind x in e) 15 | | e₀ e₁ : app 16 | | "(" ")" : unit 17 | | "(" e ")" : paren notex (expand := return e) 18 | 19 | nonterminal Environment, Γ := 20 | | "ε" : empty 21 | | Γ ", " x " : " τ : ext (id x) 22 | | Γ₀ ", " Γ₁ : append notex (expand := return .mkCApp `LottExamples.STLC.Environment.append #[Γ₀, Γ₁]) 23 | 24 | nonterminal (parent := Term) Value, v := 25 | | "λ " x ". " e : lam (bind x in e) 26 | | "(" ")" : unit 27 | 28 | end LottExamples.STLC 29 | -------------------------------------------------------------------------------- /Lott/Elab/OrJudgement.lean: -------------------------------------------------------------------------------- 1 | import Lott.Elab.Basic 2 | import Lott.Parser.OrJudgement 3 | 4 | namespace Lott 5 | 6 | open Lean 7 | open Lean.Elab 8 | 9 | @[macro Lott.judgementEmbed] 10 | private 11 | def orJudgementImpl : Macro := fun stx => do 12 | let `([[$j₀:Lott.Judgement ∨ $j₁:Lott.Judgement]]) := stx | Macro.throwUnsupported 13 | ``([[$j₀:Lott.Judgement]] ∨ [[$j₁:Lott.Judgement]]) 14 | 15 | @[lott_tex_elab orJudgement] 16 | private 17 | def orJudgementTexElab : TexElab := fun profile ref stx => do 18 | let `(Lott.Judgement| $j₀:Lott.Judgement ∨ $j₁:Lott.Judgement) := stx | throwUnsupportedSyntax 19 | let j₀Tex ← texElabSymbolOrJudgement j₀.raw.getKind profile ref j₀ 20 | let j₁Tex ← texElabSymbolOrJudgement j₁.raw.getKind profile ref j₁ 21 | return s!"{j₀Tex} \\, \\lottsym\{∨} \\, {j₁Tex}" 22 | 23 | end Lott 24 | -------------------------------------------------------------------------------- /Lott/Elab/ImpliesJudgement.lean: -------------------------------------------------------------------------------- 1 | import Lott.Elab.Basic 2 | import Lott.Parser.ImpliesJudgement 3 | 4 | namespace Lott 5 | 6 | open Lean 7 | open Lean.Elab 8 | 9 | @[macro Lott.judgementEmbed] 10 | private 11 | def impliesJudgementImpl : Macro := fun stx => do 12 | let `([[$j₀:Lott.Judgement ⇒ $j₁:Lott.Judgement]]) := stx | Macro.throwUnsupported 13 | ``([[$j₀:Lott.Judgement]] → [[$j₁:Lott.Judgement]]) 14 | 15 | @[lott_tex_elab impliesJudgement] 16 | private 17 | def impliesJudgementTexElab : TexElab := fun profile ref stx => do 18 | let `(Lott.Judgement| $j₀:Lott.Judgement ⇒ $j₁:Lott.Judgement) := stx | throwUnsupportedSyntax 19 | let j₀Tex ← texElabSymbolOrJudgement j₀.raw.getKind profile ref j₀ 20 | let j₁Tex ← texElabSymbolOrJudgement j₁.raw.getKind profile ref j₁ 21 | return s!"{j₀Tex} \\, \\lottsym\{⇒} \\, {j₁Tex}" 22 | 23 | end Lott 24 | -------------------------------------------------------------------------------- /lake-manifest.json: -------------------------------------------------------------------------------- 1 | {"version": "1.1.0", 2 | "packagesDir": ".lake/packages", 3 | "packages": 4 | [{"url": "https://github.com/leanprover-community/aesop", 5 | "type": "git", 6 | "subDir": null, 7 | "scope": "", 8 | "rev": "56a2c80b209c253e0281ac4562a92122b457dcc0", 9 | "name": "aesop", 10 | "manifestFile": "lake-manifest.json", 11 | "inputRev": "v4.17.0", 12 | "inherited": false, 13 | "configFile": "lakefile.toml"}, 14 | {"url": "https://github.com/leanprover-community/batteries", 15 | "type": "git", 16 | "subDir": null, 17 | "scope": "", 18 | "rev": "efcc7d9bd9936ecdc625baf0d033b60866565cd5", 19 | "name": "batteries", 20 | "manifestFile": "lake-manifest.json", 21 | "inputRev": "v4.17.0", 22 | "inherited": true, 23 | "configFile": "lakefile.toml"}], 24 | "name": "lott", 25 | "lakeDir": ".lake"} 26 | -------------------------------------------------------------------------------- /Lott/Tactic/Termination.lean: -------------------------------------------------------------------------------- 1 | theorem sizeOf_fst_lt_of_prod_mem [SizeOf α] {as : List (α × α)} (h : (a, b) ∈ as) : sizeOf a < sizeOf as := by 2 | apply Nat.lt_trans 3 | · show sizeOf a < sizeOf (a, b) 4 | rw [Prod.mk.sizeOf_spec] 5 | apply Nat.lt_add_right 6 | rw [Nat.add_comm] 7 | exact Nat.lt_succ_self _ 8 | · exact List.sizeOf_lt_of_mem h 9 | 10 | macro_rules 11 | | `(tactic| decreasing_trivial) => `(tactic| first 12 | | with_reducible 13 | apply Nat.lt_of_lt_of_le (sizeOf_fst_lt_of_prod_mem ?h) 14 | case h => assumption 15 | simp_arith) 16 | 17 | theorem sizeOf_snd_lt_of_prod_mem [SizeOf α] {as : List (α × α)} (h : (a, b) ∈ as) : sizeOf b < sizeOf as := by 18 | apply Nat.lt_trans 19 | · show sizeOf b < sizeOf (a, b) 20 | rw [Prod.mk.sizeOf_spec] 21 | apply Nat.lt_add_left_iff_pos.mpr 22 | rw [Nat.add_comm] 23 | exact Nat.succ_pos _ 24 | · exact List.sizeOf_lt_of_mem h 25 | 26 | macro_rules 27 | | `(tactic| decreasing_trivial) => `(tactic| first 28 | | with_reducible 29 | apply Nat.lt_of_lt_of_le (sizeOf_snd_lt_of_prod_mem ?h) 30 | case h => assumption 31 | simp_arith) 32 | -------------------------------------------------------------------------------- /nix/flake.nix: -------------------------------------------------------------------------------- 1 | { 2 | description = "lott"; 3 | 4 | inputs = { 5 | nixpkgs.follows = "lean4-nix/nixpkgs"; 6 | flake-utils.url = "github:numtide/flake-utils"; 7 | lean4-nix.url = "github:lenianiva/lean4-nix"; 8 | }; 9 | 10 | outputs = { self, nixpkgs, flake-utils, lean4-nix }: 11 | flake-utils.lib.eachDefaultSystem (system: 12 | let 13 | pkgs = import nixpkgs { 14 | overlays = [ (lean4-nix.readToolchainFile ../lean-toolchain) ]; 15 | inherit system; 16 | }; 17 | inherit (pkgs) clangStdenv mkShell texlab texliveFull; 18 | inherit (pkgs.lean) Init Lake lean-all; 19 | in 20 | { 21 | # TODO: Add overlay and package output. 22 | 23 | devShells = rec { 24 | default = (mkShell.override { stdenv = clangStdenv; }) { 25 | packages = [ lean-all ]; 26 | shellHook = '' 27 | export LEAN_SRC_PATH=${Init.src}:${Lake.src} 28 | ''; 29 | }; 30 | 31 | tex = default.overrideAttrs (oldAttrs: { 32 | nativeBuildInputs = oldAttrs.nativeBuildInputs ++ [ 33 | texlab 34 | texliveFull 35 | ]; 36 | }); 37 | }; 38 | }); 39 | } 40 | -------------------------------------------------------------------------------- /Lott/Linter/UnusedVariables.lean: -------------------------------------------------------------------------------- 1 | import Lean.Linter 2 | 3 | namespace Lean.Linter 4 | 5 | /- Term embedding syntax. -/ 6 | 7 | /- 8 | You can't put a hole inside an embedding to ignore something, so it's generally desirable to not be 9 | bugged about embedding pattern variables being unused. 10 | -/ 11 | register_option linter.unusedVariables.lottSymbolEmbeddingPatternVars : Bool := { 12 | defValue := false 13 | descr := 14 | "enable the 'unused variables' linter to mark unused lott symbol embedding pattern variables" 15 | } 16 | 17 | def getLinterUnusedVariablesLottSymbolEmbeddingPatternVars (o : Options) : Bool := 18 | o.get linter.unusedVariables.lottSymbolEmbeddingPatternVars.name 19 | (Lean.Linter.getLinterUnusedVariables o && 20 | linter.unusedVariables.lottSymbolEmbeddingPatternVars.defValue) 21 | 22 | @[unused_variables_ignore_fn] 23 | private 24 | def symbolEmbeddingPatternVars : Lean.Linter.IgnoreFunction := fun _ stack opts => 25 | !getLinterUnusedVariablesLottSymbolEmbeddingPatternVars opts && ( 26 | let stackBeforeMatchAlt := stack.takeWhile fun (stx, _) => 27 | !stx.isOfKind ``Lean.Parser.Term.matchAlt 28 | stackBeforeMatchAlt.length != stack.length && 29 | stackBeforeMatchAlt.any fun (stx, _) => stx.isOfKind `Lott.symbolEmbed 30 | ) 31 | 32 | end Lean.Linter 33 | -------------------------------------------------------------------------------- /Lott/Elab/NotExistentialJudgement.lean: -------------------------------------------------------------------------------- 1 | import Lott.Elab.Basic 2 | import Lott.Parser.NotExistentialJudgement 3 | 4 | namespace Lott 5 | 6 | open Lean 7 | open Lean.Elab 8 | 9 | @[macro judgementEmbed] 10 | private 11 | def existentialJudgementImpl : Macro 12 | | `([[∄ $binders, $«judgement»:Lott.Judgement]]) => 13 | ``(¬(∃ $binders, [[$«judgement»:Lott.Judgement]])) 14 | | _ => Macro.throwUnsupported 15 | 16 | @[lott_tex_elab notExistentialJudgement] 17 | private 18 | def notExistentialJudgementTexElab : TexElab := fun profile ref stx => do 19 | let `(notExistentialJudgement| ∄ $binders, $«judgement»:Lott.Judgement) := stx 20 | | throwUnsupportedSyntax 21 | let binderIdents ← match binders with 22 | | `(explicitBinders| $bis:binderIdent* $[: $_]?) => pure bis 23 | | `(explicitBinders| $[($bis* : $_)]*) => pure bis.flatten 24 | | _ => throwUnsupportedSyntax 25 | let binderTexs := ", ".intercalate <| Array.toList <| ← binderIdents.mapM fun 26 | | `(binderIdent| _) => return "_" 27 | | `(binderIdent| $i:ident) => return i.getId.toString false |>.texEscape 28 | | _ => throwUnsupportedSyntax 29 | -- NOTE: type is intentionally omitted. 30 | let judgementTex ← texElabSymbolOrJudgement judgement.raw.getKind profile ref «judgement» 31 | 32 | return s!"\\lottsym\{∄} \\, {binderTexs}, {judgementTex}" 33 | 34 | end Lott 35 | -------------------------------------------------------------------------------- /LottExamples/SystemF/Syntax.lean: -------------------------------------------------------------------------------- 1 | import Lott 2 | 3 | namespace LottExamples.SystemF 4 | 5 | locally_nameless 6 | metavar TypeVar, a 7 | 8 | nonterminal «Type», A, B := 9 | | a : var 10 | | A " → " B : arr 11 | | "∀ " a ". " A : forall' (bind a in A) 12 | | "(" A ")" : paren notex (expand := return A) 13 | 14 | locally_nameless 15 | metavar TermVar, x, y 16 | 17 | nonterminal Term, E, F := 18 | | x : var 19 | | "λ " x " : " A ". " E : lam (bind x in E) 20 | | E F : app 21 | | "Λ " a ". " E : typeGen (bind a in E) 22 | | E " [" A "]" : typeApp 23 | | "(" E ")" : paren notex (expand := return E) 24 | 25 | nosubst 26 | nonterminal Environment, Γ := 27 | | "ε" : empty 28 | | Γ ", " x " : " A : termVarExt (id x) 29 | | Γ ", " a : typeVarExt (id a) 30 | | Γ₀ ", " Γ₁ : append notex (expand := return .mkCApp `LottExamples.SystemF.Environment.append #[Γ₀, Γ₁]) 31 | | "(" Γ ")" : paren notex (expand := return Γ) 32 | | Γ " [" A " / " a "]" : subst notex (id a) (expand := return .mkCApp `LottExamples.SystemF.Environment.TypeVar_subst #[Γ, a, A]) 33 | 34 | nonterminal (parent := Term) Value, V := 35 | | "λ " x " : " A ". " E : lam (bind x in E) 36 | | "Λ " a ". " E : typeGen (bind a in E) 37 | 38 | end LottExamples.SystemF 39 | -------------------------------------------------------------------------------- /Lott/Data/Name.lean: -------------------------------------------------------------------------------- 1 | import Lean.Data.Name 2 | 3 | import Lott.Data.Option 4 | 5 | namespace Lean.Name 6 | 7 | def getFinal : Name → Name 8 | | anonymous => anonymous 9 | | str _ s => str anonymous s 10 | | num _ n => num anonymous n 11 | 12 | def erasePrefix? : Name → Name → Option Name 13 | | anonymous, pf => .someIf anonymous (anonymous == pf) 14 | | name@(str pre s), pf => do 15 | if name == pf then 16 | return anonymous 17 | 18 | let pre' ← pre.erasePrefix? pf 19 | return str pre' s 20 | | name@(num pre n), pf => do 21 | if name == pf then 22 | return anonymous 23 | 24 | let pre' ← pre.erasePrefix? pf 25 | return num pre' n 26 | 27 | inductive Component where 28 | | str (s : String) 29 | | num (n : Nat) 30 | deriving BEq 31 | 32 | def fromComponents' : List Name → Name 33 | | [] => anonymous 34 | | anonymous :: cs => fromComponents' cs 35 | | c@(str ..) :: cs => c ++ fromComponents' cs 36 | | c@(num ..) :: cs => c ++ fromComponents' cs 37 | 38 | def removeCommonPrefixes (n₀ n₁ : Name) : Name × Name := 39 | let cs₀ := n₀.components 40 | let cs₁ := n₁.components 41 | let sameCount := cs₀.zipWithAll (fun c₀? c₁? => c₀? == c₁?) cs₁ |>.takeWhile id |>.length 42 | let sameCount := if sameCount == cs₀.length || sameCount == cs₁.length then 43 | sameCount - 1 44 | else 45 | sameCount 46 | (fromComponents' <| cs₀.drop sameCount, fromComponents' <| cs₁.drop sameCount) 47 | 48 | end Lean.Name 49 | -------------------------------------------------------------------------------- /Lott/Parser/Filter.lean: -------------------------------------------------------------------------------- 1 | import Lott.Parser 2 | 3 | namespace Lott 4 | 5 | open Lean.Parser 6 | 7 | private 8 | def embedCloseFn : ParserFn := fun c s => 9 | if c.input.substrEq s.pos "]]" 0 2 then 10 | s.setPos <| s.pos + "]]".endPos 11 | else 12 | s.mkUnexpectedErrorAt "expected ']]'" s.pos 13 | 14 | private partial 15 | def filterParserFnAux (startPos : String.Pos) : ParserFn := fun c s => 16 | if h : c.input.atEnd s.pos then 17 | mkNodeToken `Lott.NonEmbed startPos c s 18 | else if c.input.substrEq s.pos "[[" 0 2 then 19 | let s := mkNodeToken `Lott.NonEmbed startPos c s 20 | let s := s.setPos <| s.pos + "[[".endPos 21 | let s := orelseFn 22 | (orelseFn 23 | (atomicFn (andthenFn qualifiedSymbolParser.fn embedCloseFn)) 24 | (atomicFn (andthenFn (withPosition (categoryParser `Lott.Symbol 0)).fn embedCloseFn))) 25 | (andthenFn (withPosition (categoryParser `Lott.Judgement 0)).fn embedCloseFn) c s 26 | if s.hasError then 27 | s 28 | else 29 | filterParserFnAux s.pos c s 30 | else 31 | filterParserFnAux startPos c <| s.next' c.input s.pos h 32 | 33 | def filterParserFn : ParserFn := fun c s => filterParserFnAux s.pos c s 34 | 35 | def inputParserFn : ParserFn := orelseFn 36 | (orelseFn 37 | (atomicFn (andthenFn qualifiedSymbolParser.fn whitespace)) 38 | (atomicFn (andthenFn (withPosition (categoryParser `Lott.Judgement 0)).fn whitespace))) 39 | (andthenFn (withPosition (categoryParser `Lott.Symbol 0)).fn whitespace) 40 | 41 | syntax "#filter " str (str)? : command 42 | 43 | end Lott 44 | -------------------------------------------------------------------------------- /Lott/Elab/Options.lean: -------------------------------------------------------------------------------- 1 | import Lean 2 | 3 | namespace Lott 4 | 5 | open Lean 6 | open Lean.Elab.Command 7 | 8 | register_option lott.term : Bool := { 9 | defValue := true 10 | descr := "create inductives, substitutions, free vars functions, locally closed inductives, and more" 11 | } 12 | 13 | register_option lott.tex.example.comprehensionNoTex : Bool := { 14 | defValue := true 15 | descr := "use notex option for example syntax comprehensions in tex output" 16 | } 17 | 18 | register_option lott.tex.example.singleProductionInline : Bool := { 19 | defValue := true 20 | descr := "inline example syntax of nonterminals with a single production in tex output" 21 | } 22 | 23 | register_option lott.tex.locallyNameless : Bool := { 24 | defValue := true 25 | descr := "show locally nameless operations in tex output" 26 | } 27 | 28 | register_option lott.tex.output.dir : String := { 29 | defValue := "." 30 | descr := "directory where tex output files should be saved" 31 | } 32 | 33 | register_option lott.tex.output.sourceRelative : Bool := { 34 | defValue := true 35 | descr := "when output.dir is relative, whether it should be considered relative to the source file's parent, or to the Lean process's working directory" 36 | } 37 | 38 | register_option lott.tex.output.makeDeps : Bool := { 39 | defValue := false 40 | descr := "also output .d files for make alongside tex" 41 | } 42 | 43 | def getTexOutputSome : CommandElabM Bool := return (← getOptions).contains lott.tex.output.dir.name 44 | 45 | def getTerm : CommandElabM Bool := return (← getOptions).get lott.term.name lott.term.defValue 46 | 47 | end Lott 48 | -------------------------------------------------------------------------------- /Lott/Elab/JudgementComprehension.lean: -------------------------------------------------------------------------------- 1 | import Lott.Elab.Basic 2 | import Lott.Parser.JudgementComprehension 3 | 4 | namespace Lott 5 | 6 | open Lean 7 | open Lean.Elab 8 | 9 | @[macro judgementEmbed] 10 | private 11 | def judgementComprehensionImpl : Macro := fun stx => do 12 | let `([[]]) := stx 13 | | Macro.throwUnsupported 14 | `(∀ x ∈ $c, let $p:term := x; [[$«judgement»:Lott.Judgement]]) 15 | 16 | @[lott_tex_elab judgementComprehension] 17 | private 18 | def judgementComprehensionTexElab : TexElab := fun profile ref stx => do 19 | let `(judgementComprehension| ) := stx 20 | | throwUnsupportedSyntax 21 | let judgementTex ← texElabSymbolOrJudgement «judgement».raw.getKind profile ref «judgement» 22 | if nt.isSome then 23 | if let some tex := tex? then 24 | logWarningAt tex "notex takes precedence over tex" 25 | return s!"\\lottjudgementcomprehension\{{judgementTex}}" 26 | 27 | if let some tex := tex? then 28 | return s!"\\lottjudgementcomprehensioncustom\{{judgementTex}}\{{tex}}" 29 | 30 | let some patSubstring := pat.raw.getSubstring? (withLeading := false) (withTrailing := false) 31 | | throwUnsupportedSyntax 32 | let patTex := patSubstring.toString.texEscape 33 | let some collectionSubstring := 34 | collection.raw.getSubstring? (withLeading := false) (withTrailing := false) 35 | | throwUnsupportedSyntax 36 | let collectionTex := collectionSubstring.toString.texEscape 37 | return s!"\\lottjudgementcomprehensionpatcollection\{{judgementTex}}\{{patTex}}\{{collectionTex}}" 38 | 39 | end Lott 40 | -------------------------------------------------------------------------------- /Lott/Data/Option.lean: -------------------------------------------------------------------------------- 1 | namespace Option 2 | 3 | def someIf (x : α) (b : Bool) : Option α := if b then some x else none 4 | 5 | theorem someIf_true : someIf x true = some x := rfl 6 | 7 | theorem someIf_false : someIf x false = none := rfl 8 | 9 | theorem someIf_get!_eq [Inhabited α] {x? : Option α} : someIf (get! x?) (isSome x?) = x? := by 10 | match x? with 11 | | some _ => rw [get!, isSome, someIf_true] 12 | | none => rw [isSome, someIf_false] 13 | 14 | theorem eq_of_someIf_eq_some : someIf x b = some y → x = y ∧ b = true := by 15 | intro eq 16 | rw [someIf] at eq 17 | split at eq 18 | · case isTrue h => exact ⟨some.inj eq, h⟩ 19 | · case isFalse => nomatch eq 20 | 21 | theorem eq_of_someIf_eq_none : someIf x b = none → b = false := by 22 | intro eq 23 | rw [someIf] at eq 24 | split at eq 25 | · case isTrue => nomatch eq 26 | · case isFalse h => exact Bool.not_eq_true _ |>.mp h 27 | 28 | theorem someIf_eq : someIf x b₀ = someIf y b₁ → b₀ = b₁ := by 29 | intro h 30 | match b₁ with 31 | | true => exact eq_of_someIf_eq_some h |>.right 32 | | false => exact eq_of_someIf_eq_none h 33 | 34 | theorem isSome_someIf : isSome (someIf x b) = b := by 35 | rw [someIf] 36 | split 37 | · case isTrue h => 38 | cases h 39 | rw [isSome] 40 | · case isFalse h => 41 | rw [Bool.not_eq_true _ |>.mp h, isSome] 42 | 43 | def keepIf (x : Option α) (b : Bool) : Option α := if b then x else none 44 | 45 | def mapMem (a? : Option α) (f : (a : α) → a ∈ a? → β) : Option β := match a? with 46 | | some a => f a rfl 47 | | none => none 48 | 49 | theorem mapMem_eq_map {a? : Option α} : a?.mapMem (fun a _ => f a) = a?.map f := by cases a? <;> rfl 50 | 51 | theorem map_someIf : Option.map f (someIf x b) = someIf (f x) b := by cases b <;> simp [someIf] 52 | 53 | end Option 54 | -------------------------------------------------------------------------------- /Lott/Elab/Nat.lean: -------------------------------------------------------------------------------- 1 | import Lott.Elab.Basic 2 | 3 | namespace Lott 4 | 5 | open Lean 6 | open Lean.Elab 7 | open Lean.Elab.Command 8 | open Lean.Parser 9 | 10 | declare_syntax_cat Lott.Symbol.Nat 11 | 12 | def addNatAlias (alias : Name) : CommandElabM Unit := do 13 | setEnv <| aliasExt.addEntry (← getEnv) { canon := `Nat, alias, tex? := none } 14 | let parserIdent := mkIdent <| `nat ++ (alias.appendAfter "_parser") 15 | elabCommand <| ← 16 | `(@[Lott.Symbol.Nat_parser] 17 | private 18 | def $parserIdent : Parser := 19 | leadingNode `Lott.Symbol.Nat maxPrec <| identPrefix $(quote <| alias.toString)) 20 | 21 | run_cmd addNatAlias `n 22 | 23 | @[Lott.Symbol.Nat_parser] 24 | private 25 | def nat.num_parser : Parser := leadingNode `Lott.Symbol.Nat maxPrec numLit 26 | 27 | @[Lott.Symbol.Nat_parser] 28 | private 29 | def nat.term_parser : Parser := 30 | leadingNode `Lott.Symbol.Nat maxPrec <| "(" >> checkLineEq >> termParser >> checkLineEq >> ")" 31 | 32 | @[macro symbolEmbed] 33 | private 34 | def natImpl : Macro 35 | | .node _ `Lott.symbolEmbed #[ 36 | .atom _ "[[", 37 | .node _ `Lott.Symbol.Nat #[n@(.ident ..)], 38 | .atom _ "]]" 39 | ] 40 | | .node _ `Lott.symbolEmbed #[ 41 | .atom _ "[[", 42 | .node _ `Lott.Symbol.Nat #[n@(.node _ `num _)], 43 | .atom _ "]]" 44 | ] 45 | | .node _ `Lott.symbolEmbed #[ 46 | .atom _ "[[", 47 | .node _ `Lott.Symbol.Nat #[.atom _ "(", n, .atom _ ")"], 48 | .atom _ "]]" 49 | ] => return n 50 | | _ => Macro.throwUnsupported 51 | 52 | @[lott_tex_elab Lott.Symbol.Nat] 53 | private 54 | def natTexElab : TexElab 55 | | _, _, .node _ `Lott.Symbol.Nat #[n@(.ident ..)] 56 | | _, _, .node _ `Lott.Symbol.Nat #[n@(.node _ `num _)] => texElabIdx n 57 | | _, _, .node _ `Lott.Symbol.Nat #[.atom _ "(", n, .atom _ ")"] => do 58 | let some n := n.getSubstring? (withLeading := false) (withTrailing := false) 59 | | throwUnsupportedSyntax 60 | return s!"({n.toString.texEscape})" 61 | | _, _, _ => throwUnsupportedSyntax 62 | 63 | end Lott 64 | -------------------------------------------------------------------------------- /LottExamples/main.tex: -------------------------------------------------------------------------------- 1 | \documentclass{article} 2 | 3 | \usepackage[T1]{fontenc} 4 | \usepackage[utf8]{inputenc} 5 | \usepackage{unicode-math-input} 6 | 7 | \usepackage{amsmath} 8 | \usepackage{mathpartir} 9 | \usepackage{mathtools} 10 | 11 | \newcommand{\lottalias}[1]{#1} 12 | \newcommand{\lottaliassep}{,} 13 | \newcommand{\lottproduction}[1]{#1} 14 | \newcommand{\lottkw}[1]{\textbf{#1}} 15 | \newcommand{\lottsym}[1]{{#1}} 16 | \newcommand{\lottproductionsep}{\mid} 17 | \setlength\tabcolsep{2pt} % For tabular environment spacing. 18 | \newcommand{\lottnonterminal}[3]{#1 & $#2$ & $\Coloneqq$ & $#3$} 19 | \newcommand{\lotthypothesis}[1]{#1} 20 | \newcommand{\lottinferencerule}[3]{\inferrule[#1]{#2}{#3}} 21 | \newcommand{\lottinferencerulesep}{\and} 22 | \newcommand{\lottjudgement}[3]{ 23 | \noindent\fbox{$#2$}\hfill\emph{#1} 24 | \nopagebreak 25 | \begin{mathparpagebreakable} 26 | #3 27 | \end{mathparpagebreakable} 28 | } 29 | 30 | \begin{document} 31 | 32 | \section{Lambda} 33 | 34 | \begin{tabular}{l @{\hspace{2em}} r c l} 35 | \input{Lambda/term.tex} 36 | \end{tabular} 37 | 38 | \medskip 39 | 40 | \input{Lambda/BetaReduction.tex} 41 | 42 | \input{Lambda/EtaReduction.tex} 43 | 44 | \input{Lambda/intro.tex} 45 | 46 | \section{STLC} 47 | 48 | \begin{tabular}{l @{\hspace{2em}} r c l} 49 | \input{STLC/Type.tex} \\ 50 | \input{STLC/Term.tex} \\ 51 | \input{STLC/Environment.tex} \\ 52 | \input{STLC/Value.tex} 53 | \end{tabular} 54 | 55 | \medskip 56 | 57 | \input{STLC/Environment/VarIn.tex} 58 | 59 | \input{STLC/Environment/WellFormedness.tex} 60 | 61 | \input{STLC/Typing.tex} 62 | 63 | \input{STLC/Reduction.tex} 64 | 65 | \section{System F} 66 | 67 | \begin{tabular}{l @{\hspace{2em}} r c l} 68 | \input{SystemF/Type.tex} \\ 69 | \input{SystemF/Term.tex} \\ 70 | \input{SystemF/Environment.tex} \\ 71 | \input{SystemF/Value.tex} 72 | \end{tabular} 73 | 74 | \medskip 75 | 76 | \input{SystemF/EnvironmentWellFormedness.tex} 77 | 78 | \input{SystemF/TypeVarInEnvironment.tex} 79 | 80 | \input{SystemF/TermVarInEnvironment.tex} 81 | 82 | \input{SystemF/Typing.tex} 83 | 84 | \input{SystemF/OperationalSemantics.tex} 85 | 86 | \end{document} 87 | -------------------------------------------------------------------------------- /Lott/Data/String.lean: -------------------------------------------------------------------------------- 1 | import Lott.Data.Substring 2 | 3 | namespace String 4 | 5 | private 6 | def toPascalParts (s : String) : Array String := Id.run do 7 | let rawParts := s.data.splitBy (·.isUpper == ·.isUpper) 8 | let mut parts : Array String := #[] 9 | for rawPart in rawParts do 10 | if rawPart[0]!.isLower then 11 | if let some last := parts.back? then 12 | parts := parts.pop.push <| last ++ ⟨rawPart⟩ 13 | else 14 | parts := parts.push ⟨rawPart⟩ 15 | continue 16 | 17 | if rawPart.length > 1 then 18 | parts := parts ++ #[⟨rawPart.dropLast⟩, rawPart.getLast!.toString] 19 | else 20 | parts := parts.push ⟨rawPart⟩ 21 | return parts 22 | 23 | def pascalToSnake (s : String) : String := 24 | "_".intercalate <| s.toPascalParts.toList.map String.toLower 25 | 26 | def pascalToTitle (s : String) : String := " ".intercalate s.toPascalParts.toList 27 | 28 | def texEscape (s : String) : String := 29 | join <| s.data.map fun 30 | | c@'&' | c@'%' | c@'$' | c@'#' | c@'_' | c@'{' | c@'}' => "\\" ++ c.toString 31 | | '~' => "\\textasciitilde" 32 | | '^' => "\\textasciicircum" 33 | | '\\' => "\\textbackslash" 34 | | c => c.toString 35 | 36 | /- 37 | Based on `munge` in libcpp/mkdeps.cc from the GCC source code. This is 38 | imperfect as some special characters can't be escaped. 39 | -/ 40 | def makeEscape (s : String) : String := Id.run do 41 | let mut res := "" 42 | let mut slashes := 0 43 | for c in s.data do 44 | match c with 45 | | '\\' => slashes := slashes + 1 46 | | '$' => 47 | res := res.push '$' 48 | slashes := 0 49 | | ':' => 50 | res := res.push '\\' 51 | slashes := 0 52 | | ' ' | '\t' => 53 | /- 54 | `munge`'s source contains a comment here that says: "A 55 | space or tab preceded by 2N+1 backslashes represents N 56 | backslashes followed by space..." 57 | -/ 58 | res := res.pushn '\\' <| slashes + 1 59 | slashes := 0 60 | | '#' => 61 | res := res.push '\\' 62 | slashes := 0 63 | | _ => slashes := 0 64 | res := res.push c 65 | res 66 | 67 | def dropPrefixes (s pre : String) : Substring := s.toSubstring.dropPrefixes pre.toSubstring 68 | 69 | end String 70 | -------------------------------------------------------------------------------- /Lott/Elab/UniversalJudgement.lean: -------------------------------------------------------------------------------- 1 | import Lott.Elab.Basic 2 | import Lott.Parser.UniversalJudgement 3 | 4 | namespace Lott 5 | 6 | open Lean 7 | open Lean.Elab 8 | 9 | @[macro judgementEmbed] 10 | private 11 | def universalJudgementImpl : Macro 12 | | `([[∀ $[$binders]* $[$type?]?, $«judgement»:Lott.Judgement]]) => 13 | ``(∀ $binders* $[$type?:typeSpec]?, [[$«judgement»:Lott.Judgement]]) 14 | | `([[∀ $i:ident $bp:binderPred, $«judgement»:Lott.Judgement]]) => 15 | ``(∀ $i:ident, satisfies_binder_pred% $i $bp → [[$«judgement»:Lott.Judgement]]) 16 | | `([[∀ $i:ident, $«judgement»:Lott.Judgement]]) => 17 | ``(∀ $i:ident, [[$«judgement»:Lott.Judgement]]) 18 | | _ => Macro.throwUnsupported 19 | 20 | @[lott_tex_elab universalJudgement] 21 | private 22 | def universalJudgementTexElab : TexElab := fun profile ref stx => do 23 | let `(universalJudgement| ∀ $binders* $[$type?]?, $«judgement»:Lott.Judgement) := stx 24 | | throwUnsupportedSyntax 25 | let binderTexs := ", ".intercalate <| Array.toList <| binders.map fun 26 | | `(hole| _) => "_" 27 | | `(bracketedBinder| {$is*}) => ", ".intercalate <| Array.toList <| is.map fun 28 | | `(hole| _) => "_" 29 | | `(ident| $i) => i.getId.toString false |>.texEscape 30 | | `(ident| $i) => i.getId.toString false |>.texEscape 31 | -- NOTE: type is intentionally omitted. 32 | let judgementTex ← texElabSymbolOrJudgement judgement.raw.getKind profile ref «judgement» 33 | 34 | let opts ← getOptions 35 | let locallyNameless := opts.get lott.tex.locallyNameless.name lott.tex.locallyNameless.defValue 36 | if !locallyNameless then 37 | return judgementTex 38 | 39 | return s!"\\lottsym\{∀} \\, {binderTexs}, {judgementTex}" 40 | 41 | @[lott_tex_elab universalPredJudgement] 42 | private 43 | def universalPredJudgementTexElab : TexElab 44 | | profile, ref, `(universalPredJudgement| ∀ $i:ident $bp:binderPred, $«judgement»:Lott.Judgement) => do 45 | let identTex := i.getId.toString false |>.texEscape 46 | let bpSym := bp.raw.getArg 0 |>.getAtomVal 47 | let some bpTermSubstring := 48 | bp.raw.getArg 1 |>.getSubstring? (withLeading := false) (withTrailing := false) 49 | | throwUnsupportedSyntax 50 | let bpTermTex := bpTermSubstring.toString.texEscape 51 | let judgementTex ← texElabSymbolOrJudgement judgement.raw.getKind profile ref «judgement» 52 | 53 | let opts ← getOptions 54 | let locallyNameless := opts.get lott.tex.locallyNameless.name lott.tex.locallyNameless.defValue 55 | if bpSym == "∉" && !locallyNameless then 56 | return judgementTex 57 | 58 | return s!"\\lottsym\{∀} \\, {identTex} \\, \\lottsym\{{bpSym}} \\, {bpTermTex}, {judgementTex}" 59 | | _, _, _ => throwUnsupportedSyntax 60 | 61 | end Lott 62 | -------------------------------------------------------------------------------- /LottExamples/STLC/Semantics.lean: -------------------------------------------------------------------------------- 1 | import Lott.Elab.UniversalJudgement 2 | import LottExamples.STLC.Syntax 3 | 4 | namespace LottExamples.STLC 5 | 6 | judgement_syntax x " ≠ " x' : VarId.Ne (id x, x') 7 | 8 | judgement VarId.Ne := fun (x x' : VarId) => x ≠ x' 9 | 10 | judgement_syntax x " ∈ " "fv" "(" e ")" : Term.InFreeVars (id x) 11 | 12 | judgement Term.InFreeVars := fun x (e : Term) => x ∈ e.freeVars 13 | 14 | judgement_syntax x " ∉ " "fv" "(" e ")" : Term.NotInFreeVars (id x) 15 | 16 | judgement Term.NotInFreeVars := fun x e => ¬[[x ∈ fv(e)]] 17 | 18 | namespace Environment 19 | 20 | termonly 21 | def append (Γ₀ : Environment) : Environment → Environment 22 | | [[ε]] => Γ₀ 23 | | [[Γ₁, x : τ]] => Γ₀.append Γ₁ |>.ext x τ 24 | 25 | termonly 26 | def empty_append (Γ : Environment) : Environment.empty.append Γ = Γ := match Γ with 27 | | [[ε]] => rfl 28 | | [[Γ', x : τ]] => by rw [append, Γ'.empty_append] 29 | 30 | termonly 31 | def append_assoc {Γ₀ : Environment} : Γ₀.append (Γ₁.append Γ₂) = (Γ₀.append Γ₁).append Γ₂ := by 32 | match Γ₂ with 33 | | [[ε]] => rfl 34 | | [[Γ₂', x : τ]] => rw [append, append, append_assoc, ← append] 35 | 36 | termonly 37 | def domain : Environment → List VarId 38 | | empty => [] 39 | | ext Γ x _ => x :: Γ.domain 40 | 41 | end Environment 42 | 43 | judgement_syntax x " : " τ " ∈ " Γ : Environment.VarIn (id x) 44 | 45 | judgement Environment.VarIn where 46 | 47 | ──────────────── head 48 | x : τ ∈ Γ, x : τ 49 | 50 | x : τ ∈ Γ 51 | x ≠ x' 52 | ────────────────── ext 53 | x : τ ∈ Γ, x' : τ' 54 | 55 | judgement_syntax x " ∉ " Γ : Environment.VarNotIn (id x) 56 | 57 | judgement Environment.VarNotIn := fun x Γ => ∀ τ, ¬[[x : τ ∈ Γ]] 58 | 59 | judgement_syntax x " ∈ " "dom" "(" Γ ")" : Environment.VarInDom (id x) 60 | 61 | judgement Environment.VarInDom := fun x (Γ : Environment) => x ∈ Γ.domain 62 | 63 | judgement_syntax x " ∉ " "dom" "(" Γ ")" : Environment.VarNotInDom (id x) 64 | 65 | judgement Environment.VarNotInDom := fun x Γ => ¬[[x ∈ dom(Γ)]] 66 | 67 | judgement_syntax "⊢ " Γ : Environment.WellFormedness 68 | 69 | judgement Environment.WellFormedness where 70 | 71 | ─── empty 72 | ⊢ ε 73 | 74 | ⊢ Γ 75 | x ∉ dom(Γ) 76 | ────────── ext 77 | ⊢ Γ, x : τ 78 | 79 | judgement_syntax Γ " ⊢ " e " : " τ : Typing 80 | 81 | judgement Typing where 82 | 83 | ⊢ Γ 84 | x : τ ∈ Γ 85 | ───────── var 86 | Γ ⊢ x : τ 87 | 88 | ∀ x ∉ I, Γ, x : τ₀ ⊢ e^x : τ₁ 89 | ───────────────────────────── lam {I : List VarId} 90 | Γ ⊢ λ x. e : τ₀ → τ₁ 91 | 92 | Γ ⊢ e₀ : τ₀ → τ₁ 93 | Γ ⊢ e₁ : τ₀ 94 | ──────────────── app 95 | Γ ⊢ e₀ e₁ : τ₁ 96 | 97 | ───────────── unit 98 | Γ ⊢ () : Unit 99 | 100 | judgement_syntax e " ↦ " e' : Reduction 101 | 102 | judgement Reduction where 103 | 104 | e₀ ↦ e₀' 105 | ────────────── appl 106 | e₀ e₁ ↦ e₀' e₁ 107 | 108 | e ↦ e' 109 | ────────── appr 110 | v e ↦ v e' 111 | 112 | ─────────────────── lamApp 113 | (λ x. e) v ↦ e^^v/x 114 | 115 | end LottExamples.STLC 116 | -------------------------------------------------------------------------------- /lakefile.lean: -------------------------------------------------------------------------------- 1 | import Lake 2 | import Lean.Elab.Frontend 3 | 4 | open Lake DSL 5 | 6 | def noterm := get_config? noterm |>.isSome 7 | 8 | def notex := get_config? notex |>.isSome 9 | 10 | -- TODO: Use lib leanOptions for everything instead of args once string escaping is fixed. 11 | def args := if notex then #[] else #[s!"-Dweak.lott.tex.output.dir={__dir__}"] 12 | 13 | package lott where 14 | moreGlobalServerArgs := args 15 | 16 | require aesop from git "https://github.com/leanprover-community/aesop" @ "v4.17.0" 17 | 18 | lean_lib Lott 19 | 20 | lean_lib LottExamples where 21 | leanOptions := (if noterm then #[⟨`lott.term, false⟩] else #[]) ++ 22 | if notex then #[] else #[⟨`lott.tex.output.sourceRelative, false⟩] 23 | moreLeanArgs := args 24 | 25 | open System 26 | 27 | @[implemented_by Lean.enableInitializersExecution] 28 | opaque enableInitializersExecution : IO Unit 29 | 30 | inductive Filterable where 31 | | file (input output : FilePath) 32 | | dir (path : FilePath) 33 | 34 | /-- 35 | Filter files containing nonterminal or judgement syntax 36 | 37 | USAGE: 38 | lake run lott-filter [...] 39 | -/ 40 | script «lott-filter» args do 41 | let spec :: «namespace» :: filterables := args | 42 | IO.eprintln "USAGE: lake run lott-filter [...]" 43 | return 2 44 | let currentDir ← IO.currentDir 45 | let mkAbs name := if FilePath.isAbsolute name then name else FilePath.join currentDir name 46 | let rec parseFilterables : List String → IO (Except String String) 47 | | input :: filterables => do 48 | let isDir ← FilePath.isDir input 49 | if isDir then 50 | let rest ← parseFilterables filterables 51 | return rest.map (s!"#filter {mkAbs input |>.toString.quote}\n" ++ ·) 52 | let output :: filterables := filterables 53 | | return .error "no output path following non-directory input" 54 | let rest ← parseFilterables filterables 55 | return rest.map 56 | (s!"#filter {mkAbs input |>.toString.quote} {mkAbs output |>.toString.quote}\n" ++ ·) 57 | | [] => return .ok "" 58 | match ← parseFilterables filterables with 59 | | .error err => do 60 | IO.eprintln s!"USAGE: lake run lott-filter [...]\n{err}" 61 | return 2 62 | | .ok inputCommands => 63 | let ws ← getWorkspace 64 | let filter := "Lott.Elab.Filter" 65 | match parseTargetSpecs ws [filter, spec] with 66 | | .error cliError => 67 | IO.eprintln cliError 68 | return 1 69 | | .ok specs => 70 | for spec in specs do 71 | unless spec.buildable do 72 | IO.eprintln <| CliError.invalidBuildTarget spec.info.key.toSimpleString 73 | return 1 74 | ws.runBuild <| buildSpecs specs 75 | 76 | Lean.searchPathRef.set ws.augmentedLeanPath 77 | enableInitializersExecution 78 | 79 | let input := s!"import {filter}\nimport {spec}\nnamespace {«namespace»}\n" ++ 80 | inputCommands 81 | let opts := Lean.KVMap.empty.insert `lott.tex.output.dir <| .ofString "/dev/null" 82 | let (_, ok) ← Lean.Elab.runFrontend input opts "LottFilterScript.lean" `LottFilterScript 83 | return (!ok).toUInt32 84 | -------------------------------------------------------------------------------- /nix/flake.lock: -------------------------------------------------------------------------------- 1 | { 2 | "nodes": { 3 | "flake-parts": { 4 | "inputs": { 5 | "nixpkgs-lib": "nixpkgs-lib" 6 | }, 7 | "locked": { 8 | "lastModified": 1727826117, 9 | "narHash": "sha256-K5ZLCyfO/Zj9mPFldf3iwS6oZStJcU4tSpiXTMYaaL0=", 10 | "owner": "hercules-ci", 11 | "repo": "flake-parts", 12 | "rev": "3d04084d54bedc3d6b8b736c70ef449225c361b1", 13 | "type": "github" 14 | }, 15 | "original": { 16 | "owner": "hercules-ci", 17 | "repo": "flake-parts", 18 | "type": "github" 19 | } 20 | }, 21 | "flake-utils": { 22 | "inputs": { 23 | "systems": "systems" 24 | }, 25 | "locked": { 26 | "lastModified": 1731533236, 27 | "narHash": "sha256-l0KFg5HjrsfsO/JpG+r7fRrqm12kzFHyUHqHCVpMMbI=", 28 | "owner": "numtide", 29 | "repo": "flake-utils", 30 | "rev": "11707dc2f618dd54ca8739b309ec4fc024de578b", 31 | "type": "github" 32 | }, 33 | "original": { 34 | "owner": "numtide", 35 | "repo": "flake-utils", 36 | "type": "github" 37 | } 38 | }, 39 | "lean4-nix": { 40 | "inputs": { 41 | "flake-parts": "flake-parts", 42 | "nixpkgs": "nixpkgs" 43 | }, 44 | "locked": { 45 | "lastModified": 1758037834, 46 | "narHash": "sha256-P2rBelR9lgFJPE8kEV1wwX1LHK6xWlVcQqEez84ss6E=", 47 | "owner": "lenianiva", 48 | "repo": "lean4-nix", 49 | "rev": "d87ff09bad00ca4addd471a68a968f2aa29c7c36", 50 | "type": "github" 51 | }, 52 | "original": { 53 | "owner": "lenianiva", 54 | "repo": "lean4-nix", 55 | "type": "github" 56 | } 57 | }, 58 | "nixpkgs": { 59 | "locked": { 60 | "lastModified": 1743095683, 61 | "narHash": "sha256-gWd4urRoLRe8GLVC/3rYRae1h+xfQzt09xOfb0PaHSk=", 62 | "owner": "nixos", 63 | "repo": "nixpkgs", 64 | "rev": "5e5402ecbcb27af32284d4a62553c019a3a49ea6", 65 | "type": "github" 66 | }, 67 | "original": { 68 | "owner": "nixos", 69 | "ref": "nixos-unstable", 70 | "repo": "nixpkgs", 71 | "type": "github" 72 | } 73 | }, 74 | "nixpkgs-lib": { 75 | "locked": { 76 | "lastModified": 1727825735, 77 | "narHash": "sha256-0xHYkMkeLVQAMa7gvkddbPqpxph+hDzdu1XdGPJR+Os=", 78 | "type": "tarball", 79 | "url": "https://github.com/NixOS/nixpkgs/archive/fb192fec7cc7a4c26d51779e9bab07ce6fa5597a.tar.gz" 80 | }, 81 | "original": { 82 | "type": "tarball", 83 | "url": "https://github.com/NixOS/nixpkgs/archive/fb192fec7cc7a4c26d51779e9bab07ce6fa5597a.tar.gz" 84 | } 85 | }, 86 | "root": { 87 | "inputs": { 88 | "flake-utils": "flake-utils", 89 | "lean4-nix": "lean4-nix", 90 | "nixpkgs": [ 91 | "lean4-nix", 92 | "nixpkgs" 93 | ] 94 | } 95 | }, 96 | "systems": { 97 | "locked": { 98 | "lastModified": 1681028828, 99 | "narHash": "sha256-Vy1rq5AaRuLzOxct8nz4T6wlgyUR7zLU309k9mBC768=", 100 | "owner": "nix-systems", 101 | "repo": "default", 102 | "rev": "da67096a3b9bf56a91d16901293e51ba5b49a27e", 103 | "type": "github" 104 | }, 105 | "original": { 106 | "owner": "nix-systems", 107 | "repo": "default", 108 | "type": "github" 109 | } 110 | } 111 | }, 112 | "root": "root", 113 | "version": 7 114 | } 115 | -------------------------------------------------------------------------------- /Lott/Environment/Basic.lean: -------------------------------------------------------------------------------- 1 | import Lean.Data.Trie 2 | import Lean.Environment 3 | import Lott.IR 4 | 5 | namespace Lott 6 | 7 | open Lean 8 | open Lean.Data 9 | 10 | structure Alias where 11 | canon : Name 12 | alias : Name 13 | tex? : Option String 14 | 15 | instance : Inhabited Alias where default := { canon := default, alias := default, tex? := default } 16 | 17 | structure AliasState where 18 | byAlias : Trie Alias 19 | allCanon : NameSet 20 | 21 | instance : Inhabited AliasState where default := { byAlias := default, allCanon := default } 22 | 23 | initialize aliasExt : PersistentEnvExtension Alias Alias AliasState ← 24 | registerPersistentEnvExtension { 25 | mkInitial := return default 26 | addImportedFn := fun aliasss => return { 27 | byAlias := 28 | aliasss.flatten.foldl (init := .empty) fun acc a => acc.upsert a.alias.toString fun _ => a 29 | allCanon := aliasss.flatten.map (·.canon) |> RBTree.fromArray (cmp := Name.quickCmp) 30 | } 31 | addEntryFn := fun { byAlias, allCanon } a => { 32 | byAlias := byAlias.insert a.alias.toString a 33 | allCanon := allCanon.insert a.canon 34 | } 35 | exportEntriesFn := fun { byAlias, .. } => byAlias.values 36 | } 37 | 38 | structure Symbol where 39 | qualified : Name 40 | normalProds : NameMap (Array IR) 41 | substitutions : Array (Name × Name) 42 | texPrePost? : Option (String × String) 43 | profiles : Array Name 44 | 45 | instance : Inhabited Symbol where 46 | default := { 47 | qualified := default 48 | normalProds := default 49 | substitutions := default 50 | texPrePost? := default 51 | profiles := default 52 | } 53 | 54 | abbrev SymbolState := NameMap Symbol 55 | 56 | initialize symbolExt : PersistentEnvExtension Symbol Symbol SymbolState ← 57 | registerPersistentEnvExtension { 58 | mkInitial := return default 59 | addImportedFn := fun symss => 60 | return symss.flatten.foldl (init := mkNameMap _) fun acc sym => acc.insert sym.qualified sym 61 | addEntryFn := fun st sym => st.insert sym.qualified sym 62 | exportEntriesFn := RBMap.fold (cmp := Name.quickCmp) (init := #[]) fun acc _ sym => acc.push sym 63 | } 64 | 65 | structure Judgement where 66 | name : Name 67 | ir : Array IR 68 | ids : Array Name 69 | profiles : Array Name 70 | 71 | structure JudgementState where 72 | byName : NameMap Judgement 73 | all : NameSet 74 | 75 | instance : Inhabited JudgementState where default := { byName := default, all := default } 76 | 77 | initialize judgementExt : PersistentEnvExtension Judgement Judgement JudgementState ← registerPersistentEnvExtension { 78 | mkInitial := return default 79 | addImportedFn := fun jds => return { 80 | byName := jds.flatten.foldl (init := mkNameMap _) fun acc jd => acc.insert jd.name jd 81 | all := jds.flatten.map (·.name) |> RBTree.fromArray (cmp := Name.quickCmp) 82 | } 83 | addEntryFn := fun { byName, all } jd => { 84 | byName := byName.insert jd.name jd 85 | all := all.insert jd.name 86 | } 87 | exportEntriesFn := fun { byName, .. } => 88 | byName.fold (cmp := Name.quickCmp) (init := #[]) fun acc _ jd => acc.push jd 89 | } 90 | 91 | structure Child where 92 | canon : Name 93 | parent : Name 94 | 95 | abbrev ChildState := NameMap Child 96 | 97 | initialize childExt : PersistentEnvExtension Child Child ChildState ← registerPersistentEnvExtension { 98 | mkInitial := return default 99 | addImportedFn := fun children => 100 | return children.flatten.foldl (init := mkNameMap _) fun acc child => acc.insert child.canon child 101 | addEntryFn := fun st child => st.insert child.canon child 102 | exportEntriesFn := RBMap.fold (cmp := Name.quickCmp) (init := #[]) fun acc _ child => acc.push child 103 | } 104 | 105 | end Lott 106 | -------------------------------------------------------------------------------- /Lott/Elab/Filter.lean: -------------------------------------------------------------------------------- 1 | import Lott.Elab.Basic 2 | import Lott.Parser.Filter 3 | 4 | namespace Lott 5 | 6 | open Lean 7 | open Lean.Parser 8 | open Lean.Elab.Command 9 | open System FilePath 10 | open IO.FS 11 | 12 | private 13 | def mkAbsolute (name : String) : CommandElabM FilePath := do 14 | if isAbsolute name then 15 | return name 16 | let some parent := parent <| ← getFileName | 17 | throwError "failed to resolve parent of current file" 18 | return parent / name 19 | 20 | elab_rules : command 21 | | `(#filter $inputName:str $[$outputName?:str]?) => do 22 | if !(← getTexOutputSome) then 23 | throwError "can't #filter when when lott.tex.output.dir is unset" 24 | 25 | let inputPath ← mkAbsolute inputName.getString 26 | if ← isDir inputPath then 27 | if let some outputName := outputName? then 28 | throwErrorAt outputName "#filter output cannot be specified when input is a directory" 29 | 30 | let mut failed := false 31 | 32 | for entry in ← readDir inputPath do 33 | if extension entry.path != some "lottinput" then 34 | continue 35 | 36 | let input ← String.trimRight <$> readFile entry.path 37 | 38 | let s := mkParserState input 39 | let ictx := mkInputContext input entry.path.toString 40 | let env ← getEnv 41 | let parserFn := orelseFn 42 | (orelseFn 43 | (atomicFn (andthenFn qualifiedSymbolParser.fn whitespace)) 44 | (atomicFn (andthenFn (withPosition (categoryParser `Lott.Judgement 0)).fn whitespace))) 45 | (andthenFn (withPosition (categoryParser `Lott.Symbol 0)).fn whitespace) 46 | let s := parserFn.run ictx { env, options := {} } (getTokenTable env) s 47 | if !s.allErrors.isEmpty then 48 | logWarning s!"parse failure for `{input}` in {entry.fileName}: {s.toErrorMsg ictx}" 49 | failed := true 50 | continue 51 | else if !ictx.input.atEnd s.pos then 52 | let msg := s.mkError "end of input" |>.toErrorMsg ictx 53 | logWarning s!"parse failure for `{input}` in {entry.fileName}: {msg}" 54 | failed := true 55 | continue 56 | 57 | let mut stx := s.stxStack.back 58 | if stx.getKind == choiceKind then 59 | let possibilities := ", ".intercalate <| stx.getArgs.map (·.getKind.toString) |>.toList 60 | logWarning s!"`{input}` from {entry.fileName} is ambiguous between: {possibilities}" 61 | failed := true 62 | continue 63 | if stx.getKind == `Lott.QualifiedSymbol then 64 | stx := stx.getArg 2 65 | 66 | let output ← liftTermElabM <| texElabSymbolOrJudgement stx.getKind default inputName stx 67 | writeFile (withExtension entry.path "tex") <| output ++ "\n" 68 | 69 | if failed then 70 | throwError "some inputs were not processed succesfully" 71 | return 72 | 73 | let outputPath ← if let some outputName := outputName? then 74 | mkAbsolute outputName.getString 75 | else 76 | let some outputName := outputName?.map TSyntax.getString |>.orElse 77 | fun () => inputName.getString.dropSuffix? ".lotttmpl" |>.map Substring.toString 78 | | throwErrorAt inputName 79 | "#filter input name must end with '.lotttmpl' if output name is omitted" 80 | 81 | mkAbsolute outputName 82 | 83 | let input ← readFile inputPath 84 | 85 | let s := mkParserState input 86 | let ictx := mkInputContext input inputName.getString 87 | let env ← getEnv 88 | let s := filterParserFn.run ictx { env, options := {} } (getTokenTable env) s 89 | if !s.allErrors.isEmpty then 90 | throwError s.toErrorMsg ictx 91 | else if !ictx.input.atEnd s.pos then 92 | throwError s.mkError "end of input" |>.toErrorMsg ictx 93 | 94 | let output ← s.stxStack.toSubarray.toArray.filterMapM fun 95 | | .node _ `Lott.NonEmbed #[.atom _ s] => return s 96 | | stx => do 97 | let s ← liftTermElabM <| texElabSymbolOrJudgement stx.getKind default inputName stx 98 | return "$" ++ s ++ "$" 99 | 100 | -- TODO: Make the output path here relative if possible. Needs diff_paths-like. 101 | writeMakeDeps outputPath (extraDeps := [inputPath]) 102 | writeFile outputPath <| String.join output.toList 103 | 104 | end Lott 105 | -------------------------------------------------------------------------------- /Lott/Elab/MetaVar.lean: -------------------------------------------------------------------------------- 1 | import Lott.Elab.Basic 2 | 3 | namespace Lott 4 | 5 | open Lean 6 | open Lean.Elab 7 | open Lean.Elab.Command 8 | open Lean.Parser 9 | open Lean.Parser.Term 10 | 11 | @[macro Lott.symbolEmbed] 12 | private 13 | def metavarImpl : Macro := fun 14 | | .node _ ``Lott.symbolEmbed #[ 15 | .atom _ "[[", 16 | .node _ _ #[ident@(.ident ..)], 17 | .atom _ "]]" 18 | ] => return ident 19 | | .node _ ``Lott.symbolEmbed #[ 20 | .atom _ "[[", 21 | .node _ _ #[«fun», .atom _ "@", idx], 22 | .atom _ "]]" 23 | ] => `((↑[[$(.mk «fun»):Lott.Symbol]] : _ → _) $(.mk idx):term) 24 | | .node _ ``Lott.symbolEmbed #[ 25 | .atom _ "[[", 26 | .node _ _ #[.ident .., .atom _ "$", num@(.node _ `num _)], 27 | .atom _ "]]" 28 | ] => `(.bound $(.mk num)) 29 | | _ => Macro.throwUnsupported 30 | 31 | elab_rules : command | `($[locally_nameless%$ln]? metavar $[(tex pre := $pre?, post := $post?)]? $names,*) => do 32 | let names := names.getElems.toList 33 | let canon :: aliases := names | throwUnsupportedSyntax 34 | let ns ← getCurrNamespace 35 | let qualified := ns ++ canon.getId 36 | let locallyNameless := ln.isSome 37 | 38 | -- Declare type and aliases. 39 | -- TODO: Is there any way we can make the ids have an opaque type which reveals nothing other than 40 | -- that they have a decidable equality relation? 41 | if ← getTerm then 42 | if locallyNameless then 43 | let idIdent := mkIdentFrom canon <| canon.getId.appendAfter "Id" 44 | elabCommand <| ← `(def $idIdent := Nat) 45 | elabCommand <| ← `(instance (x y : $idIdent) : Decidable (x = y) := Nat.decEq x y) 46 | elabCommand <| ← 47 | `(inductive $canon where 48 | | $(mkIdent `free):ident (id : $idIdent) 49 | | $(mkIdent `bound):ident (idx : Nat)) 50 | elabCommand <| ← `(instance : SizeOf $canon := instSizeOfDefault $canon) 51 | elabCommand <| ← 52 | `(instance (x y : $canon) : Decidable (x = y) := match x, y with 53 | | .free idx, .free idy => if h : idx = idy then 54 | isTrue <| $(mkIdent <| canon.getId ++ `free.injEq) idx idy |>.mpr h 55 | else 56 | isFalse (h <| $(mkIdent <| canon.getId ++ `free.inj) ·) 57 | | .bound idxx, .bound idxy => if h : idxx = idxy then 58 | isTrue <| $(mkIdent <| canon.getId ++ `bound.injEq) idxx idxy |>.mpr h 59 | else 60 | isFalse (h <| $(mkIdent <| canon.getId ++ `bound.inj) ·) 61 | | .free _, .bound _ 62 | | .bound _, .free _ => isFalse nofun) 63 | elabCommand <| ← `(instance : Coe $idIdent $canon where coe := .free) 64 | else 65 | elabCommand <| ← `(def $canon := Nat) 66 | elabCommand <| ← `(instance (x y : $canon) : Decidable (x = y) := Nat.decEq x y) 67 | else 68 | elabCommand <| ← `(opaque $canon : Type) 69 | 70 | -- Update environment extensions. 71 | for name in names do 72 | setEnv <| aliasExt.addEntry (← getEnv) 73 | { canon := qualified, alias := ns ++ name.getId, tex? := none } 74 | setEnv <| metaVarExt.addEntry (← getEnv) (qualified, locallyNameless) 75 | let texPrePost? ← match pre?, post? with 76 | | some pre, some post => pure <| some (pre.getString, post.getString) 77 | | none, none => pure none 78 | | _, _ => unreachable! 79 | setEnv <| symbolExt.addEntry (← getEnv) 80 | { qualified, normalProds := .empty, substitutions := #[], texPrePost?, profiles := #[] } 81 | 82 | -- Declare syntax category. For metavariables we just declare the alias name parsers directly in 83 | -- the syntax category. This differs from variable parsers for non-terminals, for which we declare 84 | -- a separate variable category, then add a category parser for the variable category to the main 85 | -- non-terminal symbol category. 86 | let catName := symbolPrefix ++ qualified 87 | let parserAttrName := catName.appendAfter "_parser" 88 | setEnv <| ← registerParserCategory (← getEnv) parserAttrName catName .symbol 89 | 90 | -- Declare parsers in category. 91 | for alias in aliases do 92 | let aliasName@(.str .anonymous nameStr) := alias.getId | throwUnsupportedSyntax 93 | let parserIdent := mkIdentFrom alias <| canon.getId ++ aliasName.appendAfter "_parser" 94 | elabCommand <| ← 95 | `(@[Lott.Symbol_parser, $(mkIdent parserAttrName):ident] 96 | private 97 | def $parserIdent : Parser := 98 | leadingNode $(quote catName) Parser.maxPrec <| Lott.identPrefix $(quote nameStr)) 99 | 100 | if locallyNameless then 101 | let parserIdent := mkIdentFrom alias <| canon.getId ++ aliasName.appendAfter "_bound_parser" 102 | elabCommand <| ← 103 | `(@[Lott.Symbol_parser, $(mkIdent parserAttrName):ident] 104 | private 105 | def $parserIdent : Parser := 106 | leadingNode $(quote catName) Parser.maxPrec <| 107 | Lott.identPrefix $(quote nameStr) >> "$" >> num) 108 | 109 | let parserIdent := mkIdentFrom canon <| canon.getId.appendAfter "_idx_parser" 110 | elabCommand <| ← 111 | `(@[Lott.Symbol_parser, $(mkIdent parserAttrName):ident] 112 | private 113 | def $parserIdent : TrailingParser := 114 | trailingNode $(quote catName) Parser.maxPrec 0 <| 115 | checkNoWsBefore >> "@" >> checkLineEq >> (Parser.ident <|> Parser.numLit)) 116 | 117 | end Lott 118 | -------------------------------------------------------------------------------- /LottExamples/SystemF/Semantics.lean: -------------------------------------------------------------------------------- 1 | import Lott.Data.List 2 | import Lott.Elab.UniversalJudgement 3 | import LottExamples.SystemF.Syntax 4 | 5 | namespace LottExamples.SystemF 6 | 7 | namespace Environment 8 | 9 | termonly 10 | def append (Γ : Environment) : Environment → Environment 11 | | empty => Γ 12 | | termVarExt Γ' x A => Γ.append Γ' |>.termVarExt x A 13 | | typeVarExt Γ' a => Γ.append Γ' |>.typeVarExt a 14 | 15 | termonly 16 | def TypeVar_subst (Γ : Environment) (a : TypeVarId) (A : «Type») := match Γ with 17 | | [[ε]] => empty 18 | | [[Γ', x : A']] => Γ'.TypeVar_subst a A |>.termVarExt x <| A'.TypeVar_subst a A 19 | | [[Γ', a']] => Γ'.TypeVar_subst a A |>.typeVarExt a' 20 | 21 | termonly 22 | def termVar_count : Environment → Nat 23 | | [[ε]] => 0 24 | | [[Γ, x : A]] => 1 + Γ.termVar_count 25 | | [[Γ, a]] => Γ.termVar_count 26 | 27 | termonly 28 | def typeVar_count : Environment → Nat 29 | | [[ε]] => 0 30 | | [[Γ, x : A]] => Γ.typeVar_count 31 | | [[Γ, a]] => 1 + Γ.typeVar_count 32 | 33 | termonly 34 | def termVar_domain : Environment → List TermVarId 35 | | [[ε]] => [] 36 | | [[Γ, x : A]] => x :: Γ.termVar_domain 37 | | [[Γ, a]] => Γ.termVar_domain 38 | 39 | termonly 40 | def typeVar_domain : Environment → List TypeVarId 41 | | [[ε]] => [] 42 | | [[Γ, x : A]] => Γ.typeVar_domain 43 | | [[Γ, a]] => a :: Γ.typeVar_domain 44 | 45 | end Environment 46 | 47 | judgement_syntax a " ≠ " a' : TypeVarNe (id a, a') 48 | 49 | judgement TypeVarNe := Ne (α := TypeVarId) 50 | 51 | judgement_syntax a " ∈ " Γ : TypeVarInEnvironment (id a) 52 | 53 | judgement TypeVarInEnvironment where 54 | 55 | ──────── head 56 | a ∈ Γ, a 57 | 58 | a ∈ Γ 59 | ──────────── termVarExt 60 | a ∈ Γ, x : A 61 | 62 | a ∈ Γ 63 | a ≠ a' 64 | ───────── typeVarExt 65 | a ∈ Γ, a' 66 | 67 | judgement_syntax a " ∉ " Γ : TypeVarNotInEnvironment (id a) 68 | 69 | judgement TypeVarNotInEnvironment := fun a Γ => ¬[[a ∈ Γ]] 70 | 71 | judgement_syntax x " ≠ " y : TermVarNe (id x, y) 72 | 73 | judgement TermVarNe := Ne (α := TermVarId) 74 | 75 | judgement_syntax x " : " A " ∈ " Γ : TermVarInEnvironment (id x) 76 | 77 | judgement TermVarInEnvironment where 78 | 79 | ──────────────── head 80 | x : A ∈ Γ, x : A 81 | 82 | x : A ∈ Γ 83 | x ≠ x' 84 | ───────────────── termVarExt 85 | x : A ∈ Γ, x' : B 86 | 87 | x : A ∈ Γ 88 | ──────────── typeVarExt 89 | x : A ∈ Γ, a 90 | 91 | judgement_syntax x " ∉ " Γ : TermVarNotInEnvironment (id x) 92 | 93 | judgement TermVarNotInEnvironment := fun x Γ => ∀ A, ¬[[x : A ∈ Γ]] 94 | 95 | judgement_syntax a " ∈ " "ftv" "(" A ")" : Type.InFreeTypeVars (id a) 96 | 97 | judgement Type.InFreeTypeVars := fun a (A : «Type») => a ∈ A.freeTypeVars 98 | 99 | judgement_syntax a " ∉ " "ftv" "(" A ")" : Type.NotInFreeTypeVars (id a) 100 | 101 | judgement Type.NotInFreeTypeVars := fun a A => ¬[[a ∈ ftv(A)]] 102 | 103 | judgement_syntax x " ∈ " "fv" "(" E ")" : Term.InFreeTermVars (id x) 104 | 105 | judgement Term.InFreeTermVars := fun x (E : Term) => x ∈ E.freeTermVars 106 | 107 | judgement_syntax x " ∉ " "fv" "(" E ")" : Term.NotInFreeTermVars (id x) 108 | 109 | judgement Term.NotInFreeTermVars := fun x E => ¬[[x ∈ fv(E)]] 110 | 111 | judgement_syntax a " ∈ " "ftv" "(" E ")" : Term.InFreeTypeVars (id a) 112 | 113 | judgement Term.InFreeTypeVars := fun a (E : Term) => a ∈ E.freeTypeVars 114 | 115 | judgement_syntax a " ∉ " "ftv" "(" E ")" : Term.NotInFreeTypeVars (id a) 116 | 117 | judgement Term.NotInFreeTypeVars := fun a E => ¬[[a ∈ ftv(E)]] 118 | 119 | judgement_syntax Γ " ⊢ " A : TypeWellFormedness 120 | 121 | judgement TypeWellFormedness where 122 | 123 | a ∈ Γ 124 | ───── var 125 | Γ ⊢ a 126 | 127 | Γ ⊢ A 128 | Γ ⊢ B 129 | ───────── arr 130 | Γ ⊢ A → B 131 | 132 | ∀ a ∉ I, Γ, a ⊢ A^a 133 | ────────────────────────────── forall' {I : List TypeVarId} 134 | Γ ⊢ ∀ a. A 135 | 136 | judgement_syntax x " ∈ " "dom" "(" Γ ")" : TermVarInEnvironmentDomain (id x) 137 | 138 | judgement TermVarInEnvironmentDomain := fun x (Γ : Environment) => x ∈ Γ.termVar_domain 139 | 140 | judgement_syntax x " ∉ " "dom" "(" Γ ")" : TermVarNotInEnvironmentDomain (id x) 141 | 142 | judgement TermVarNotInEnvironmentDomain := fun x Γ => ¬[[x ∈ dom(Γ)]] 143 | 144 | judgement_syntax a " ∈ " "dom" "(" Γ ")" : TypeVarInEnvironmentDomain (id a) 145 | 146 | judgement TypeVarInEnvironmentDomain := fun a (Γ : Environment) => a ∈ Γ.typeVar_domain 147 | 148 | judgement_syntax a " ∉ " "dom" "(" Γ ")" : TypeVarNotInEnvironmentDomain (id a) 149 | 150 | judgement TypeVarNotInEnvironmentDomain := fun a Γ => ¬[[a ∈ dom(Γ)]] 151 | 152 | judgement_syntax "⊢ " Γ : EnvironmentWellFormedness 153 | 154 | judgement EnvironmentWellFormedness where 155 | 156 | ─── empty 157 | ⊢ ε 158 | 159 | ⊢ Γ 160 | Γ ⊢ A 161 | x ∉ dom(Γ) 162 | ────────── termVarExt 163 | ⊢ Γ, x : A 164 | 165 | ⊢ Γ 166 | a ∉ dom(Γ) 167 | ────────── typeVarExt 168 | ⊢ Γ, a 169 | 170 | judgement_syntax Γ " ⊢ " E " : " A : Typing 171 | 172 | judgement Typing where 173 | 174 | ⊢ Γ 175 | x : A ∈ Γ 176 | ───────── var 177 | Γ ⊢ x : A 178 | 179 | Γ ⊢ A 180 | ∀ x ∉ I, Γ, x : A ⊢ E^x : B 181 | ────────────────────────────────────── lam {I : List TermVarId} 182 | Γ ⊢ λ x : A. E : A → B 183 | 184 | Γ ⊢ E : A → B 185 | Γ ⊢ F : A 186 | ───────────── app 187 | Γ ⊢ E F : B 188 | 189 | ∀ a ∉ I, Γ, a ⊢ E^a : A^a 190 | ──────────────────────────────────── typeGen {I : List TypeVarId} 191 | Γ ⊢ Λ a. E : ∀ a. A 192 | 193 | Γ ⊢ E : ∀ a. A 194 | Γ ⊢ B 195 | ────────────────── typeApp 196 | Γ ⊢ E [B] : A^^B/a 197 | 198 | judgement_syntax E " -> " F : OperationalSemantics (tex := s!"{E} \\, \\lottsym\{→} \\, {F}") 199 | 200 | judgement OperationalSemantics where 201 | 202 | E -> E' 203 | ─────────── appl 204 | E F -> E' F 205 | 206 | F -> F' 207 | ─────────── appr 208 | V F -> V F' 209 | 210 | ──────────────────────── lamApp 211 | (λ x : A. E) V -> E^^V/x 212 | 213 | E -> E' 214 | ─────────────── typeApp 215 | E [A] -> E' [A] 216 | 217 | ────────────────────── typeGenApp 218 | (Λ a. E) [A] -> E^^A/a 219 | 220 | end LottExamples.SystemF 221 | -------------------------------------------------------------------------------- /Lott/Parser.lean: -------------------------------------------------------------------------------- 1 | import Lean.Elab 2 | import Lean.Parser 3 | 4 | namespace Lott 5 | 6 | open Lean Parser 7 | open Lean.Parser.Syntax 8 | 9 | /- Common trailing syntax. -/ 10 | 11 | private 12 | def identPrefixFn («prefix» : String) : ParserFn := fun c s => 13 | let s := tokenFn ["identifier"] c s 14 | if s.hasError then 15 | s 16 | else 17 | match s.stxStack.back with 18 | | .ident _ _ val _ => 19 | if !prefix.isPrefixOf val.getRoot.getString! then 20 | s.mkUnexpectedTokenError s!"identifier beginning with '{«prefix»}'" 21 | else 22 | s 23 | | _ => s.mkUnexpectedTokenError "identifier" 24 | 25 | def identPrefix («prefix» : String) : Parser where 26 | fn := identPrefixFn «prefix» 27 | info := mkAtomicInfo "ident" 28 | 29 | @[combinator_parenthesizer identPrefix] 30 | def identPrefix.parenthesizer (_ : String) := 31 | PrettyPrinter.Parenthesizer.visitToken 32 | 33 | @[combinator_formatter identPrefix] 34 | def identPrefix.formatter (_ : String) := 35 | PrettyPrinter.Formatter.rawIdentNoAntiquot.formatter 36 | 37 | /- Metavariable syntax. -/ 38 | 39 | def texPrePostConfig := 40 | "(" >> nonReservedSymbol "tex " >> nonReservedSymbol "pre" >> " := " >> strLit >> ", " >> 41 | nonReservedSymbol "post" >> " := " >> strLit >> ") " 42 | 43 | syntax "locally_nameless "? "metavar " (texPrePostConfig)? ident,+ : command 44 | 45 | /- Non-terminal syntax. -/ 46 | 47 | declare_syntax_cat Lott.Symbol 48 | declare_syntax_cat Lott.Production 49 | declare_syntax_cat Lott.NonTerminal 50 | 51 | def prodArg := leading_parser 52 | Parser.optional (atomic (ident >> checkNoWsBefore "no space before ':'" >> ":")) >> syntaxParser argPrec 53 | 54 | def bindConfig := 55 | " (" >> nonReservedSymbol "bind " >> ident >> optional (" in " >> sepBy1 ident ", ") >> ")" 56 | 57 | def idConfig := " (" >> nonReservedSymbol "id " >> sepBy1 ident ", " >> ")" 58 | 59 | def expandConfig := " (" >> nonReservedSymbol "expand" >> " := " >> termParser >> ")" 60 | 61 | def texConfig := " (" >> nonReservedSymbol "tex" >> optional ident >> " := " >> termParser >> ")" 62 | 63 | def strLitTexConfig := " (" >> nonReservedSymbol "tex" >> " := " >> strLit >> ")" 64 | 65 | syntax ppLine "|" (ppSpace prodArg)+ " : " withPosition(ident (lineEq " nosubst")? " notex"?) atomic(bindConfig)? atomic(idConfig)? atomic(expandConfig)? (texConfig)* : Lott.Production 66 | 67 | private 68 | def parent := nonReservedSymbol "parent" 69 | 70 | syntax "nosubst "? "nonterminal " atomic("(" parent " := " ident ") ")? (texPrePostConfig)? (ident " notex"? (strLitTexConfig)?),+ " := " Lott.Production* : Lott.NonTerminal 71 | 72 | syntax Lott.NonTerminal : command 73 | 74 | /- Judgement syntax. -/ 75 | 76 | declare_syntax_cat Lott.Judgement 77 | declare_syntax_cat Lott.InferenceRuleUpper 78 | declare_syntax_cat Lott.InferenceRule 79 | declare_syntax_cat Lott.JudgementDeclRHS 80 | declare_syntax_cat Lott.JudgementDecl 81 | 82 | syntax "judgement_syntax" (ppSpace stx)+ " : " ident atomic(idConfig)? (texConfig)* : command 83 | 84 | private 85 | def bracketedBinder := Term.bracketedBinder 86 | 87 | syntax Lott.Judgement : Lott.InferenceRuleUpper 88 | 89 | syntax ident " := " Lott.Symbol : Lott.InferenceRuleUpper 90 | 91 | syntax "notex " ("for " ident)? Lott.InferenceRuleUpper : Lott.InferenceRuleUpper 92 | 93 | def commentConfig := 94 | " (" >> nonReservedSymbol "comment" >> optional ident >> optional (" := " >> strLit) >> ")" 95 | 96 | syntax withPosition(ppLine Lott.InferenceRuleUpper)* ppLine "─"+ withPosition(ident atomic(lineEq commentConfig)* (lineEq bracketedBinder)* (lineEq " notex")?) withPosition(ppLine Lott.Judgement) : Lott.InferenceRule 97 | 98 | syntax " where" ppLine Lott.InferenceRule* : Lott.JudgementDeclRHS 99 | 100 | syntax " := " term : Lott.JudgementDeclRHS 101 | 102 | syntax "judgement " ident Lott.JudgementDeclRHS : Lott.JudgementDecl 103 | 104 | syntax Lott.JudgementDecl : command 105 | 106 | syntax (name := zetaReduce) "zeta_reduce% " term : term 107 | 108 | /- Conditional syntax. -/ 109 | 110 | syntax "termonly " command : command 111 | 112 | /- Term embedding syntax. -/ 113 | 114 | syntax (name := symbolEmbed) "[[" withPosition(Lott.Symbol) "]]" : term 115 | 116 | syntax (name := judgementEmbed) "[[" withPosition(Lott.Judgement) "]]" : term 117 | 118 | /- 119 | Like Lean/Parser/Extension.lean, Lean/PrettyPrinter/Formatter.lean, and 120 | Lean/PrettyPrinter/Parenthesizer.lean, but with the Lott.Symbol prefix automatically added. 121 | -/ 122 | private 123 | def parserOfStackFn (offset : Nat) : ParserFn := fun c s => Id.run do 124 | let stack := s.stxStack 125 | if stack.size < offset + 1 then 126 | return s.mkUnexpectedError "failed to determine parser using syntax stack, stack is too small" 127 | let .ident _ _ parserName _ := stack.get! (stack.size - offset - 1) 128 | | s.mkUnexpectedError s!"failed to determine parser using syntax stack, the specified element on the stack is not an identifier: {stack.get! (stack.size - offset)} {stack.get! (stack.size - offset - 1)} {stack.get! (stack.size - offset - 2)} {stack.get! (stack.size - offset - 3)}" 129 | let parserName := mkIdent <| `Lott.Symbol ++ parserName 130 | let oldStackSize := s.stackSize 131 | let s ← match c.resolveParserName ⟨parserName⟩ with 132 | | [.category cat] => categoryParserFn cat c s 133 | | [_] => return s.mkUnexpectedError s!"parser name {parserName} should be a lott symbol category, not a parser or alias" 134 | | _ :: _ :: _ => return s.mkUnexpectedError s!"ambiguous parser name {parserName}" 135 | | [] => return s.mkUnexpectedError s!"unknown parser {parserName}" 136 | if !s.hasError && s.stackSize != oldStackSize + 1 then 137 | return s.mkUnexpectedError "expected parser to return exactly one syntax object" 138 | s 139 | 140 | private 141 | def parserOfStack (offset : Nat) (prec : Nat := 0) : Parser where 142 | fn := adaptCacheableContextFn ({ · with prec }) (parserOfStackFn offset) 143 | 144 | open PrettyPrinter Formatter in 145 | @[combinator_formatter parserOfStack] 146 | private 147 | def parserOfStack.formatter (offset : Nat) (_prec : Nat := 0) : Formatter := do 148 | let st ← get 149 | let stx := st.stxTrav.parents.back!.getArg (st.stxTrav.idxs.back! - offset) 150 | formatterForKind stx.getKind 151 | 152 | open PrettyPrinter Parenthesizer in 153 | @[combinator_parenthesizer parserOfStack] 154 | private 155 | def parserOfStack.parenthesizer (offset : Nat) (_prec : Nat := 0) : Parenthesizer := do 156 | let st ← get 157 | let stx := st.stxTrav.parents.back!.getArg (st.stxTrav.idxs.back! - offset) 158 | parenthesizerForKind stx.getKind 159 | 160 | def qualifiedSymbolParser := leadingNode `Lott.QualifiedSymbol Parser.maxPrec <| 161 | ident >> "| " >> incQuotDepth (parserOfStack 1) 162 | 163 | @[term_parser] 164 | def qualifiedSymbolEmbed := leading_parser 165 | "[[" >> qualifiedSymbolParser >> "]]" 166 | 167 | end Lott 168 | -------------------------------------------------------------------------------- /Lott/IR.lean: -------------------------------------------------------------------------------- 1 | import Lean 2 | import Lott.Data.Option 3 | import Lott.Environment.MetaVar 4 | import Lott.Parser 5 | 6 | namespace Lott 7 | 8 | open Lean 9 | open Lean.Elab 10 | open Lean.Elab.Command 11 | open Lean.Syntax 12 | open Lean.Parser 13 | open Lean.Parser.Command 14 | open Lean.Parser.Term 15 | 16 | def optionalPrefix := `Lott.Optional 17 | 18 | def symbolPrefix := `Lott.Symbol 19 | 20 | def sepByPrefix := `Lott.SepBy 21 | 22 | def variablePrefix := `Lott.Variable 23 | 24 | -- Useful for avoiding "un-uniqueification" resulting from the use of `eraseMacroScopes`. 25 | def _root_.Lean.Name.obfuscateMacroScopes (n : Name) : Name := 26 | let scopesView := extractMacroScopes n 27 | scopesView.scopes.foldl (.str · <| toString ·) scopesView.name 28 | 29 | -- TODO: Is there a way to use Lean's existing parser resolution instead of this custom stuff? 30 | 31 | mutual 32 | private 33 | inductive IRType where 34 | | category (n : Name) 35 | | atom (s : String) 36 | | sepBy (ir : Array IR) (sep : String) 37 | | optional (ir : Array IR) 38 | deriving Inhabited, BEq 39 | 40 | inductive IR where 41 | | mk (l : Ident) (t : IRType) 42 | deriving Inhabited, BEq 43 | end 44 | 45 | namespace IR 46 | 47 | partial 48 | def containsName (ir : Array IR) (name : Name) : Bool := 49 | ir.any fun (mk l ir) => 50 | l.getId == name || 51 | match ir with 52 | | .sepBy ir _ | .optional ir => containsName ir name 53 | | _ => false 54 | 55 | mutual 56 | 57 | private partial 58 | def toParser' (canon : Name) : IR → CommandElabM (Option Term) 59 | | mk _ (.category n) => ``(categoryParser $(quote <| symbolPrefix ++ n) Parser.maxPrec) 60 | | mk _ (.atom s) => if s == "" then return none else ``(symbol $(mkStrLit s)) 61 | | mk l (.sepBy ir sep) => do 62 | let canon' := canon ++ l.getId |>.obfuscateMacroScopes 63 | let catName := sepByPrefix ++ (← getCurrNamespace) ++ canon' 64 | let parserAttrName := catName.appendAfter "_parser" 65 | 66 | setEnv <| ← Parser.registerParserCategory (← getEnv) parserAttrName catName .symbol 67 | 68 | let attrIdent := mkIdent parserAttrName 69 | let (val, type) ← toParser ir canon' sepByPrefix 70 | if type != (← `(term| Parser)) then throwError "invalid left recursive sepBy syntax" 71 | let parserIdent := mkIdentFrom l <| canon'.appendAfter "_parser" 72 | elabCommand <| ← 73 | `(@[$attrIdent:ident] 74 | private 75 | def $parserIdent : Parser := $val) 76 | 77 | let comprehensionIdent := mkIdentFrom l <| canon'.appendAfter "_comprehension_parser" 78 | elabCommand <| ← 79 | `(@[$attrIdent:ident] 80 | private 81 | def $comprehensionIdent : Parser := 82 | leadingNode $(quote catName) Parser.maxPrec <| 83 | "> withPosition (categoryParser $(quote catName) 0) >> " // " >> 84 | optional (atomic strLitTexConfig) >> termParser >> " in " >> termParser >> 85 | optional " notex" >> " />") 86 | 87 | let sepIdent := mkIdentFrom l <| canon'.appendAfter "_sep_parser" 88 | elabCommand <| ← 89 | `(@[$attrIdent:ident] 90 | private 91 | def $sepIdent : TrailingParser := 92 | trailingNode $(quote catName) Parser.maxPrec 0 <| 93 | checkLineEq >> $(quote sep) >> categoryParser $(quote catName) 0) 94 | 95 | ``(Parser.optional (categoryParser $(quote catName) 0)) 96 | | mk l (.optional ir) => do 97 | let canon' := canon ++ l.getId |>.obfuscateMacroScopes 98 | let catName := optionalPrefix ++ (← getCurrNamespace) ++ canon' 99 | let parserAttrName := catName.appendAfter "_parser" 100 | 101 | setEnv <| ← Parser.registerParserCategory (← getEnv) parserAttrName catName .symbol 102 | 103 | let attrIdent := mkIdent parserAttrName 104 | let (val, type) ← toParser ir canon' optionalPrefix 105 | if type != (← `(term| Parser)) then throwError "invalid left recursive optional syntax" 106 | let parserIdent := mkIdentFrom l <| canon'.appendAfter "_parser" 107 | elabCommand <| ← 108 | `(@[$attrIdent:ident] 109 | private 110 | def $parserIdent : Parser := $val) 111 | 112 | let comprehensionIdent := mkIdentFrom l <| canon'.appendAfter "_comprehension_parser" 113 | elabCommand <| ← 114 | `(@[$attrIdent:ident] 115 | private 116 | def $comprehensionIdent : Parser := 117 | leadingNode $(quote catName) Parser.maxPrec <| 118 | "> withPosition (categoryParser $(quote catName) 0) >> " // " >> 119 | optional (atomic strLitTexConfig) >> termParser >> optional " notex" >> " />") 120 | 121 | ``(Parser.optional (categoryParser $(quote catName) 0)) 122 | 123 | private partial 124 | def toParserSeq (canon : Name) (ir : Array IR) : CommandElabM Term := do 125 | let parserTerms ← ir.filterMapM <| toParser' canon 126 | 127 | if parserTerms.size == 0 then 128 | throwUnsupportedSyntax 129 | 130 | parserTerms.foldlM (start := 1) (init := parserTerms[0]!) 131 | fun acc t => do ``($acc >> checkLineEq >> $t) 132 | 133 | partial 134 | def toParser (ir : Array IR) (canon catPrefix : Name) : CommandElabM (Term × Term) := do 135 | let qualified := (← getCurrNamespace) ++ canon 136 | let catName := catPrefix ++ qualified 137 | if let some (mk _ (.category n)) := ir[0]? then 138 | if n == qualified then 139 | let rest ← toParserSeq canon <| ir.extract 1 ir.size 140 | return ( 141 | ← ``(trailingNode $(quote catName) Parser.maxPrec 0 <| checkLineEq >> $rest), 142 | ← ``(TrailingParser), 143 | ) 144 | 145 | if let some (mk _ (.optional _)) := ir[0]? then 146 | throwError "unsupported optionally left recursive syntax" 147 | 148 | let rest ← toParserSeq canon ir 149 | return (← ``(leadingNode $(quote catName) Parser.maxPrec $rest), ← ``(Parser)) 150 | 151 | end 152 | 153 | mutual 154 | partial 155 | def toType (ids binders : Array Name) : IR → CommandElabM (Option Term) 156 | | IR.mk l (.category n) => do 157 | for binder in binders do 158 | if l.getId == binder && (metaVarExt.getState (← getEnv)).find! n then 159 | return none 160 | for id in ids do 161 | if l.getId == id && (metaVarExt.getState (← getEnv)).find! n then 162 | return some <| mkIdent <| n.appendAfter "Id" 163 | return some <| mkIdent n 164 | | IR.mk _ (.atom _) => return none 165 | | IR.mk _ (.sepBy ir _) => do 166 | let some type' ← toTypeProdSeq ids binders ir | return none 167 | ``(List $type') 168 | | IR.mk _ (.optional ir) => do 169 | let some type' ← toTypeProdSeq ids binders ir | return none 170 | return some <| ← ``(Option $type') 171 | 172 | partial 173 | def toTypeProdSeq (ids binders : Array Name) (ir : Array IR) : CommandElabM (Option Term) := do 174 | let types ← ir.filterMapM <| IR.toType ids binders 175 | let some (type' : Term) := types[0]? | return none 176 | return some <| ← types.foldlM (start := 1) (init := type') fun acc t => ``($acc × $t) 177 | end 178 | 179 | def foldrArrow (args : Array Term) (init : Term) : CommandElabM Term := 180 | args.foldrM (init := init) fun arg acc => ``($arg → $acc) 181 | 182 | def toTypeArrSeq (ir : Array IR) (init : Term) (ids binders : Array Name) : CommandElabM Term := do 183 | (← ir.filterMapM <| IR.toType ids binders) |> foldrArrow (init := init) 184 | 185 | private 186 | def toPatternArg : IR → CommandElabM (Option Term) 187 | | mk l (.category n) => `($l@(Lean.Syntax.node _ $(quote <| symbolPrefix ++ n) _)) 188 | | mk l (.atom s) => if s == "" then return none else `($l@(Lean.Syntax.atom _ $(quote s.trim))) 189 | | mk l (.sepBy ..) => `($l@(Lean.Syntax.node _ `null _)) 190 | | mk l (.optional ..) => `($l@(Lean.Syntax.node _ `null _)) 191 | 192 | def toPatternArgs (ir : Array IR) : CommandElabM (TSepArray `term ",") := 193 | ir.filterMapM IR.toPatternArg 194 | 195 | def toJoinArgs (ir : Array IR) : TSepArray `term "," := Id.run do 196 | let (elems, _) := ir.foldl (init := (#[], false)) 197 | fun | (acc, lastWasNonAtom), mk l ir => 198 | let (nonAtom, emptyAtom) := match ir with 199 | | .atom s => (false, s == "") 200 | | _ => (true, false) 201 | ( 202 | if emptyAtom then acc else 203 | acc ++ (Option.someIf (Syntax.mkStrLit "\\,") (lastWasNonAtom && nonAtom)).toArray 204 | |>.push l, 205 | nonAtom, 206 | ) 207 | let mut elemsAndSeps := #[] 208 | for (elem, i) in elems.zipIdx do 209 | if i != 0 then 210 | elemsAndSeps := elemsAndSeps.push <| mkAtom "," 211 | elemsAndSeps := elemsAndSeps.push elem 212 | return .mk elemsAndSeps 213 | 214 | def foldrProd (as : Array Term) : MacroM Term := if let some a := as.back? then 215 | as.foldrM (init := a) (start := as.size - 1) fun a acc => `(($a, $acc)) 216 | else 217 | ``(()) 218 | 219 | def foldlAnd (as : Array Term) : CommandElabM Term := if let some a := as[0]? then 220 | as.foldlM (init := a) (start := 1) fun acc a => `($acc ∧ $a) 221 | else 222 | ``(True) 223 | 224 | def foldlAppend (as : Array Term) : CommandElabM Term := if let some a := as[0]? then 225 | as.foldlM (init := a) (start := 1) fun acc a => `($acc ++ $a) 226 | else 227 | ``([]) 228 | 229 | def toTerms (as : TSyntaxArray `Lean.binderIdent) : CommandElabM (Array Term) := 230 | as.mapM fun 231 | | `(Lean.binderIdent| $h:hole) => `(term| $h:hole) 232 | | `(Lean.binderIdent| $i:ident) => `(term| $i:ident) 233 | | _ => throwUnsupportedSyntax 234 | 235 | end IR 236 | 237 | end Lott 238 | -------------------------------------------------------------------------------- /Lott/Data/List.lean: -------------------------------------------------------------------------------- 1 | import Lott.Data.Function 2 | import Lott.Data.Nat 3 | 4 | namespace List 5 | 6 | /- 7 | As suggested by Joachim Breitner on Zulip: https://leanprover.zulipchat.com/#narrow/channel/113488-general/topic/Internalizing.20congruence.20lemmas/near/398694086 8 | -/ 9 | @[specialize] 10 | def mapMem (as : List α) (f : (a : α) → a ∈ as → β) : List β := match as with 11 | | [] => [] 12 | | a :: as' => f a (.head _) :: as'.mapMem (f · <| ·.tail _) 13 | 14 | @[specialize] 15 | def mapMemIdx (as : List α) (f : Nat → (a : α) → a ∈ as → β) : List β := match as with 16 | | [] => [] 17 | | a :: as' => f as'.length a (.head _) :: as'.mapMemIdx (f · · <| ·.tail _) 18 | 19 | @[specialize] 20 | def mapMemM [Monad m] (as : List α) (f : (a : α) → a ∈ as → m β) : m (List β) := match as with 21 | | [] => return [] 22 | | a :: as' => return (← f a (.head _)) :: (← as'.mapMemM (f · <| ·.tail _)) 23 | 24 | @[specialize] 25 | def mapMemIdxM [Monad m] (as : List α) (f : Nat → (a : α) → a ∈ as → m β) : m (List β) := 26 | match as with 27 | | [] => return [] 28 | | a :: as' => return (← f as'.length a (.head _)) :: (← as'.mapMemIdxM (f · · <| ·.tail _)) 29 | 30 | @[simp] 31 | theorem mapMem_eq_map {as : List α} : as.mapMem (fun a _ => f a) = as.map f := by 32 | match as with 33 | | [] => rfl 34 | | a :: as' => rw [mapMem, map, as'.mapMem_eq_map] 35 | 36 | theorem map_singleton_flatten (xs : List α) : (xs.map fun x => [f x]).flatten = xs.map f := 37 | match xs with 38 | | [] => rfl 39 | | x :: xs' => by rw [List.map, List.map, List.flatten, List.singleton_append, xs'.map_singleton_flatten] 40 | 41 | theorem not_mem_append' {xs ys : List α} : z ∉ xs ++ ys ↔ z ∉ xs ∧ z ∉ ys where 42 | mp zninxsys := ⟨(zninxsys <| mem_append.mpr <| .inl ·), (zninxsys <| mem_append.mpr <| .inr ·)⟩ 43 | mpr | ⟨xninxs, xninys⟩, xinxsys => match mem_append.mp xinxsys with 44 | | .inl xinxs => xninxs xinxs 45 | | .inr xinys => xninys xinys 46 | 47 | theorem not_mem_cons : x ∉ y :: xs ↔ x ≠ y ∧ x ∉ xs where 48 | mp nmem := ⟨ne_of_not_mem_cons nmem, not_mem_of_not_mem_cons nmem⟩ 49 | mpr | ⟨ne, nmem⟩ => not_mem_cons_of_ne_of_not_mem ne nmem 50 | 51 | theorem not_mem_flatten : x ∉ flatten xss ↔ ∀ xs ∈ xss, x ∉ xs where 52 | mp xnmem xs xsin := by 53 | let _ :: xss' := xss 54 | rw [flatten] at xnmem 55 | match xsin with 56 | | .head _ => exact List.not_mem_append'.mp xnmem |>.left 57 | | .tail _ xsin' => 58 | apply not_mem_flatten.mp <| List.not_mem_append'.mp xnmem |>.right 59 | exact xsin' 60 | mpr nmem_of_mem := by 61 | match xss with 62 | | [] => 63 | rw [List.flatten_nil] 64 | intro xin 65 | nomatch xin 66 | | xs :: xss' => 67 | rw [List.flatten] 68 | exact List.not_mem_append'.mpr ⟨ 69 | nmem_of_mem _ <| .head _, 70 | not_mem_flatten.mpr fun _ mem => nmem_of_mem _ <| .tail _ mem 71 | ⟩ 72 | 73 | theorem exists_gt (xs : List Nat) 74 | : ∃ n : Nat, ∀ m : Nat, m ∈ xs → m < n := match xs with 75 | | [] => .intro 0 nofun 76 | | x :: xs' => 77 | let ⟨nxs', nxs'gt⟩ := xs'.exists_gt 78 | let n := max (x + 1) nxs' 79 | .intro n fun x' x'inxs => match List.mem_cons.mp x'inxs with 80 | | .inl h => by 81 | rw [h] 82 | dsimp only [n] 83 | rw [Nat.max_def] 84 | split 85 | · case isTrue h => exact Nat.lt_of_succ_le h 86 | · case isFalse h => exact Nat.lt_succ_self _ 87 | | .inr x'inxs' => Nat.lt_of_lt_of_le (nxs'gt x' x'inxs') <| Nat.le_max_right _ _ 88 | 89 | theorem exists_fresh (xs : List Nat) : ∃ n, n ∉ xs := 90 | let ⟨nxs, nxsgt⟩ := xs.exists_gt 91 | .intro nxs fun nxsinxs => Nat.not_le_of_lt (nxsgt _ nxsinxs) <| Nat.le_refl _ 92 | 93 | theorem exists_fresh_inj (xs : List Nat) : ∃ n : Nat → Nat, n.Injective' ∧ ∀ i, n i ∉ xs := 94 | let ⟨nxs, nxsgt⟩ := xs.exists_gt 95 | .intro (fun i => nxs + i) ⟨ 96 | fun _ _ eq => Nat.add_left_cancel eq, 97 | fun _ nxsinxs => Nat.not_le_of_lt (nxsgt _ nxsinxs) <| Nat.le_add_right _ _ 98 | ⟩ 99 | 100 | theorem le_sum_of_mem' {as : List Nat} (h : a ∈ as) : a ≤ as.sum := by 101 | match h with 102 | | .head _ => 103 | rw [sum_cons] 104 | exact Nat.le_add_right _ _ 105 | | .tail _ h' => 106 | rw [sum_cons] 107 | exact Nat.le_trans (Nat.le_add_left ..) <| Nat.add_le_add_iff_left.mpr <| le_sum_of_mem' h' 108 | 109 | theorem not_mem_singleton : a ∉ [b] ↔ a ≠ b := 110 | ⟨(· <| mem_singleton.mpr ·), (· <| mem_singleton.mp ·)⟩ 111 | 112 | theorem eq_of_map_eq_map_of_inj {α β : Type} {f : α → β} {l₀ l₁ : List α} 113 | (eq : List.map f l₀ = List.map f l₁) (finj : ∀ x ∈ l₀, ∀ y ∈ l₁, f x = f y → x = y) 114 | : l₀ = l₁ := by 115 | match l₁ with 116 | | [] => 117 | rw [map_nil] at eq 118 | rw [map_eq_nil_iff.mp eq] 119 | | x :: l₁' => 120 | rw [map_cons] at eq 121 | let ⟨_, _, eq₀, eq₁, eq'⟩ := map_eq_cons_iff.mp eq 122 | cases eq₀ 123 | rw [finj _ (.head _) _ (.head _) eq₁] 124 | rw [eq_of_map_eq_map_of_inj eq' fun _ mem₀ _ mem₁ => finj _ (.tail _ mem₀) _ (.tail _ mem₁)] 125 | 126 | theorem get!_mem [Inhabited α] {as : List α} (h : i < as.length) : as.get! i ∈ as := by 127 | rw [get!_eq_getD, getD_eq_getElem?_getD, getElem?_def, dif_pos h, Option.getD] 128 | exact getElem_mem _ 129 | 130 | theorem two_le_count_of_get!_eq_of_ne [BEq α] [LawfulBEq α] [Inhabited α] {as : List α} 131 | (iltlen : i < as.length) (jltlen : j < as.length) (eq : as.get! i = as.get! j) (ne : i ≠ j) 132 | : 2 ≤ List.count (as.get! i) as := by match as with 133 | | [] => 134 | rw [length_nil] at iltlen 135 | nomatch Nat.not_lt_zero _ iltlen 136 | | a :: as' => 137 | rw [count_cons] 138 | split 139 | · case isTrue h => 140 | by_cases i = 0 141 | · case pos h => 142 | cases h 143 | rw [eq] 144 | let ⟨k, eq'⟩ := Nat.exists_eq_succ_of_ne_zero ne.symm 145 | cases eq' 146 | rw [get!_cons_succ] 147 | show 1 + 1 ≤ _ 148 | rw [length_cons] at jltlen 149 | apply Nat.add_le_add_right <| Nat.succ_le_of_lt <| count_pos_iff.mpr <| get!_mem <| 150 | Nat.add_lt_add_iff_right.mp jltlen 151 | · case neg h => 152 | let ⟨k, eq'⟩ := Nat.exists_eq_succ_of_ne_zero h 153 | cases eq' 154 | rw [get!_cons_succ] 155 | by_cases j = 0 156 | · case pos h' => 157 | show 1 + 1 ≤ _ 158 | rw [length_cons] at jltlen 159 | apply Nat.add_le_add_right <| Nat.succ_le_of_lt <| count_pos_iff.mpr <| get!_mem <| 160 | Nat.add_lt_add_iff_right.mp iltlen 161 | · case neg h' => 162 | apply Nat.le_add_right_of_le 163 | let ⟨l, eq''⟩ := Nat.exists_eq_succ_of_ne_zero h' 164 | cases eq'' 165 | exact two_le_count_of_get!_eq_of_ne (Nat.add_lt_add_iff_right.mp iltlen) 166 | (Nat.add_lt_add_iff_right.mp jltlen) eq (ne <| Nat.succ_inj'.mpr ·) 167 | · case isFalse h => 168 | by_cases i = 0 169 | · case pos h' => 170 | cases h' 171 | rw [get!_cons_zero, BEq.refl] at h 172 | nomatch h 173 | · case neg h' => 174 | by_cases j = 0 175 | · case pos h'' => 176 | cases h'' 177 | rw [eq, get!_cons_zero, BEq.refl] at h 178 | nomatch h 179 | · case neg h'' => 180 | let ⟨k, eq'⟩ := Nat.exists_eq_succ_of_ne_zero h' 181 | cases eq' 182 | let ⟨k', eq''⟩ := Nat.exists_eq_succ_of_ne_zero h'' 183 | cases eq'' 184 | rw [get!_cons_succ, Nat.add_zero] 185 | exact two_le_count_of_get!_eq_of_ne (Nat.add_lt_add_iff_right.mp iltlen) 186 | (Nat.add_lt_add_iff_right.mp jltlen) eq (ne <| Nat.succ_inj'.mpr ·) 187 | 188 | theorem map_eq_id_of_eq_id_of_mem (id_of_mem : ∀ a ∈ as, f a = a) : List.map f as = as := by 189 | match as with 190 | | [] => rw [List.map_nil] 191 | | a :: as' => 192 | rw [List.map_cons, id_of_mem _ <| .head _, 193 | map_eq_id_of_eq_id_of_mem fun _ mem => id_of_mem _ <| .tail _ mem] 194 | 195 | theorem sizeOf_map_eq_of_eq_id_of_mem [SizeOf α] {f : α → α} 196 | (sizeOf_eq_of_mem : ∀ a ∈ as, sizeOf (f a) = sizeOf a) : sizeOf (List.map f as) = sizeOf as := by 197 | match as with 198 | | [] => rw [List.map_nil] 199 | | a :: as' => 200 | rw [List.map_cons, List.cons.sizeOf_spec, List.cons.sizeOf_spec, sizeOf_eq_of_mem _ <| .head _, 201 | sizeOf_map_eq_of_eq_id_of_mem fun _ mem => sizeOf_eq_of_mem _ <| .tail _ mem] 202 | 203 | theorem sum_map_eq_of_eq_of_mem {f g : α → Nat} (eq_of_mem : ∀ a ∈ as, f a = g a) 204 | : (List.map f as).sum = (as.map g).sum := by match as with 205 | | [] => rw [List.map_nil, List.map_nil] 206 | | a :: as' => 207 | rw [List.map, List.map, List.sum_cons, List.sum_cons, eq_of_mem _ <| .head _, 208 | sum_map_eq_of_eq_of_mem fun _ mem => eq_of_mem _ <| .tail _ mem] 209 | 210 | theorem get!_zip [Inhabited α] [Inhabited β] {l₁ : List α} {l₂ : List β} 211 | (length_eq : l₁.length = l₂.length) (ilt : i < l₁.length) 212 | : (zip l₁ l₂).get! i = (l₁.get! i, l₂.get! i) := by match l₁, l₂ with 213 | | [], _ => nomatch ilt 214 | | _, [] => 215 | rw [length_eq] at ilt 216 | nomatch ilt 217 | | a :: l₁', b :: l₂' => 218 | rw [zip_cons_cons] 219 | match i with 220 | | 0 => rw [get!_cons_zero, get!_cons_zero, get!_cons_zero] 221 | | i' + 1 => 222 | rw [get!_cons_succ, get!_cons_succ, get!_cons_succ] 223 | rw [length, length] at length_eq 224 | rw [length] at ilt 225 | exact get!_zip (Nat.add_one_inj.mp length_eq) <| Nat.lt_of_succ_lt_succ ilt 226 | 227 | theorem cofinite_skolem {p : Nat → α → Prop} {I : List Nat} 228 | : (∀ x ∉ I, ∃ a, p x a) → ∃ f : Nat → α, ∀ x ∉ I, p x (f x) := by 229 | intro h 230 | let ⟨x', x'nmem⟩ := I.exists_fresh 231 | let ⟨a', _⟩ := h x' x'nmem 232 | apply (Classical.skolem (b := fun _ => α) (p := fun x a => x ∉ I → p x a)).mp 233 | intro x 234 | by_cases x ∉ I 235 | · case pos xnmem => 236 | let ⟨y, h'⟩ := h _ xnmem 237 | exact ⟨y, fun _ => h'⟩ 238 | · case neg xmem => exact ⟨a', fun xnmem => nomatch xmem xnmem⟩ 239 | 240 | theorem sizeOf_append [SizeOf α] {l₁ l₂ : List α} 241 | : sizeOf (l₁ ++ l₂) + 1 = sizeOf l₁ + sizeOf l₂ := by 242 | induction l₁ <;> simp_arith 243 | assumption 244 | 245 | end List 246 | -------------------------------------------------------------------------------- /Lott/Elab/Judgement.lean: -------------------------------------------------------------------------------- 1 | import Lott.Elab.Basic 2 | import Lott.Parser 3 | 4 | namespace Lott 5 | 6 | open Lean 7 | open Lean.Elab 8 | open Lean.Elab.Command 9 | open Lean.Elab.Term 10 | open Lean.Meta 11 | open Lean.Parser 12 | open Lean.Parser.Command 13 | open Lean.Parser.Term 14 | 15 | private 16 | def judgementPrefix := `Lott.Judgement 17 | 18 | elab_rules : command | `(judgement_syntax $[$ps]* : $name $[(id $ids?,*)]? $[(tex $[$texProfile?s]? := $tex?s)]*) => do 19 | -- Declare syntax category. 20 | let ns ← getCurrNamespace 21 | let qualified := ns ++ name.getId 22 | let catName := judgementPrefix ++ qualified 23 | let parserAttrName := catName.appendAfter "_parser" 24 | setEnv <| ← Parser.registerParserCategory (← getEnv) parserAttrName catName .symbol 25 | 26 | -- Add to environment extension. 27 | let ir ← ps.mapM IR.ofStx 28 | let ids := ids?.getD (.mk #[]) |>.getElems |>.map TSyntax.getId 29 | let profiles := Std.HashSet.toArray <| Std.HashSet.ofArray <| 30 | texProfile?s.filterMap <| Option.map TSyntax.getId 31 | setEnv <| judgementExt.addEntry (← getEnv) { name := qualified, ir, ids, profiles } 32 | 33 | -- Declare parser. 34 | let (val, type) ← IR.toParser ir name.getId judgementPrefix 35 | if type != (← `(term| Parser)) then throwError "invalid left recursive judgement syntax" 36 | let parserIdent := mkIdentFrom name <| name.getId.appendAfter "_parser" 37 | elabCommand <| ← 38 | `(@[Lott.Judgement_parser, $(mkIdent parserAttrName):ident] 39 | private 40 | def $parserIdent : Parser := $val) 41 | 42 | -- Define macro and tex elab. 43 | let patternArgs ← IR.toPatternArgs ir 44 | 45 | if ← getTerm then 46 | let macroSeqItems ← IR.toMacroSeqItems ir name.getId ids #[] 47 | let exprArgs ← IR.toExprArgs ir ids #[] 48 | 49 | let macroName := mkIdentFrom name <| name.getId.appendAfter "Impl" 50 | elabCommand <| ← 51 | `(@[macro $(mkIdent `Lott.judgementEmbed)] 52 | private 53 | def $macroName : Macro := fun stx => do 54 | let Lean.Syntax.node _ ``Lott.judgementEmbed #[ 55 | Lean.Syntax.atom _ "[[", 56 | Lean.Syntax.node _ $(quote catName) #[$patternArgs,*], 57 | Lean.Syntax.atom _ "]]" 58 | ] := stx | Macro.throwUnsupported 59 | $macroSeqItems* 60 | return Syntax.mkApp (mkIdent $(quote qualified)) #[$exprArgs,*]) 61 | 62 | if ← getTexOutputSome then 63 | let catIdent := mkIdent catName 64 | let texSeqItems ← IR.toTexSeqItems ir name.getId 65 | let (alts, defaults) := texProfile?s.zip tex?s |>.partition fun (profile?, _) => profile?.isSome 66 | 67 | let profileAlternatives ← alts.mapM fun (profile?, tex) => do 68 | `(doSeqItem| if profile == $(quote profile?.get!.getId) then return $tex) 69 | 70 | let rest ← match defaults with 71 | | ⟨[]⟩ => do 72 | let joinArgs := IR.toJoinArgs ir 73 | `(term| " ".intercalate [$joinArgs,*]) 74 | | ⟨[(_, tex)]⟩ => pure tex 75 | | _ => throwUnsupportedSyntax 76 | 77 | let texElabName := mkIdentFrom name <| name.getId.appendAfter "TexElab" 78 | elabCommand <| ← 79 | `(@[lott_tex_elab $catIdent] 80 | private 81 | def $texElabName : TexElab := fun profile ref stx => do 82 | let Lean.Syntax.node _ $(quote catName) #[$patternArgs,*] := stx 83 | | throwUnsupportedSyntax 84 | $texSeqItems* 85 | $profileAlternatives* 86 | return $rest) 87 | 88 | private 89 | inductive RulesOrTerm 90 | | rules (r : TSyntaxArray `Lott.InferenceRule) 91 | | term (t : Term) 92 | 93 | private 94 | def elabJudgementDecls (jds : Array Syntax) : CommandElabM Unit := do 95 | let ns ← getCurrNamespace 96 | let jds ← jds.mapM fun jd => do 97 | let `(JudgementDecl| judgement $name $rhs) := jd | throwUnsupportedSyntax 98 | let (name, rulesOrTerm) ← match rhs with 99 | | `(JudgementDeclRHS| where $r:Lott.InferenceRule*) => pure (name, RulesOrTerm.rules r) 100 | | `(JudgementDeclRHS| := $t:term) => pure (name, .term t) 101 | | _ => throwUnsupportedSyntax 102 | let some jd := judgementExt.getState (← getEnv) |>.byName.find? <| ns ++ name.getId 103 | | throwError "judgement_syntax for {name} not given before use in judgement" 104 | return (name, jd, rulesOrTerm) 105 | 106 | if ← getTerm then 107 | elabMutualCommands <| ← jds.mapM fun (name, { ir, ids, .. }, rulesOrTerm) => do 108 | match rulesOrTerm with 109 | | .rules r => 110 | let qualified := ns ++ name.getId 111 | let type ← IR.toTypeArrSeq ir (← `(term| Prop)) ids #[] 112 | let catName := judgementPrefix ++ qualified 113 | let ctors ← r.mapM fun rule => do 114 | let `(InferenceRule| $upper:Lott.InferenceRuleUpper* $[─]* $name $[(comment $[$_]? $[:= $_]?)]* $binders* $[notex%$nt?]? $conclusion:Lott.Judgement) := rule 115 | | throwUnsupportedSyntax 116 | let conclusionKind := conclusion.raw.getKind 117 | if conclusionKind != catName then 118 | throwErrorAt conclusion 119 | "found conclusion judgement syntax kind{indentD conclusionKind}\ 120 | expected to find kind{indentD catName}\ 121 | all conclusions of inference rules in a judgement declaration must be the judgement which is being defined" 122 | 123 | let ctorType ← upper.foldrM (init := ← `(term| [[$conclusion:Lott.Judgement]])) 124 | fun 125 | | `(InferenceRuleUpper| notex $«judgement»:Lott.Judgement), acc => 126 | `([[$«judgement»:Lott.Judgement]] → $acc) 127 | | `(InferenceRuleUpper| notex $i:ident := $sym), acc => 128 | `(let $i := [[$sym:Lott.Symbol]]; $acc) 129 | | `(InferenceRuleUpper| notex for $_ $«judgement»:Lott.Judgement), acc => 130 | `([[$«judgement»:Lott.Judgement]] → $acc) 131 | | `(InferenceRuleUpper| notex for $_:ident $i:ident := $sym), acc => 132 | `(let $i := [[$sym:Lott.Symbol]]; $acc) 133 | | `(InferenceRuleUpper| $«judgement»:Lott.Judgement), acc => 134 | `([[$«judgement»:Lott.Judgement]] → $acc) 135 | | `(InferenceRuleUpper| $i:ident := $sym), acc => 136 | `(let $i := [[$sym:Lott.Symbol]]; $acc) 137 | | _, _ => throwUnsupportedSyntax 138 | `(ctor| | $name:ident $binders* : zeta_reduce% $ctorType) 139 | 140 | `(inductive $name : $type where $ctors*) 141 | | .term t => `(def $name := $t) 142 | else 143 | for (name, _) in jds do 144 | elabCommand <| ← `(opaque $name : Type) 145 | 146 | if ← getTexOutputSome then 147 | for (name, { ir, profiles, .. }, rulesOrTerm) in jds do 148 | let .rules r := rulesOrTerm | continue 149 | 150 | for profile in profiles.push default do 151 | let qualified := ns ++ name.getId 152 | let catName := judgementPrefix ++ qualified 153 | writeTexOutput (qualified ++ profile) do 154 | let nameTex := name.getId.getFinal.getString!.pascalToTitle.texEscape 155 | let exampleStx := mkNode catName <| ← toExampleSyntax ir qualified profile 156 | let syntaxTex ← liftTermElabM <| texElabSymbolOrJudgement catName profile name exampleStx 157 | let inferenceRuleTexs ← r.filterMapM fun rule => do 158 | let `(InferenceRule| $upper:Lott.InferenceRuleUpper* $[─]* $name $[(comment $[$commentProfile?s]? $[:= $comment?s]?)]* $_* $[notex%$nt?]? $conclusion:Lott.Judgement) := rule 159 | | throwUnsupportedSyntax 160 | 161 | if nt?.isSome then 162 | return none 163 | 164 | let nameTex := name.getId.getFinal.getString!.texEscape 165 | let hypothesesTexs ← upper.filterMapM fun 166 | | `(InferenceRuleUpper| notex $_:Lott.InferenceRuleUpper) => return none 167 | | `(InferenceRuleUpper| notex for $profile' $upper:Lott.InferenceRuleUpper) => do 168 | if profile == profile'.getId then 169 | return none 170 | match upper with 171 | | `(InferenceRuleUpper| $i:ident := $sym) => do 172 | let catName := sym.raw.getKind 173 | let env ← getEnv 174 | let mut idTex := i.getId.toString false |>.texEscape 175 | if let some qualified := catName.erasePrefix? symbolPrefix then 176 | if let some { texPrePost? := some (texPre, texPost), .. } := 177 | symbolExt.getState env |>.find? qualified then 178 | idTex := s!"{texPre} {idTex} {texPost}" 179 | let symTex ← liftTermElabM <| texElabSymbolOrJudgement catName profile sym sym 180 | return s!"\n\\lottlet\{{idTex}}\{{symTex}}" 181 | | `(InferenceRuleUpper| $hyp:Lott.Judgement) => do 182 | let hypTex ← liftTermElabM <| 183 | texElabSymbolOrJudgement hyp.raw.getKind profile hyp hyp 184 | return s!"\n\\lotthypothesis\{{hypTex}}" 185 | | _ => throwUnsupportedSyntax 186 | | `(InferenceRuleUpper| $i:ident := $sym) => do 187 | let catName := sym.raw.getKind 188 | let env ← getEnv 189 | let mut idTex := i.getId.toString false |>.texEscape 190 | if let some qualified := catName.erasePrefix? symbolPrefix then 191 | if let some { texPrePost? := some (texPre, texPost), .. } := 192 | symbolExt.getState env |>.find? qualified then 193 | idTex := s!"{texPre} {idTex} {texPost}" 194 | let symTex ← liftTermElabM <| texElabSymbolOrJudgement catName profile sym sym 195 | return s!"\n\\lottlet\{{idTex}}\{{symTex}}" 196 | | `(InferenceRuleUpper| $hyp:Lott.Judgement) => do 197 | let hypTex ← liftTermElabM <| 198 | texElabSymbolOrJudgement hyp.raw.getKind profile hyp hyp 199 | return s!"\n\\lotthypothesis\{{hypTex}}" 200 | | _ => throwUnsupportedSyntax 201 | let mut hypothesesTex := "\\\\".intercalate hypothesesTexs.toList 202 | let conclusionTex ← liftTermElabM <| 203 | texElabSymbolOrJudgement catName profile conclusion conclusion 204 | let (alts, defaults) := commentProfile?s.zip comment?s |>.partition fun (profile?, _) => 205 | profile?.isSome 206 | 207 | let default? ← match defaults with 208 | | ⟨[]⟩ => pure none 209 | | ⟨[(_, default)]⟩ => pure default 210 | | ⟨_⟩ => throwUnsupportedSyntax 211 | 212 | let comment? := Option.getD (dflt := default?) <| Option.map Prod.snd <| 213 | alts.find? fun (p?, _) => p?.map TSyntax.getId == some profile 214 | 215 | if let some comment := comment? then 216 | return s!"\\lottinferencerulecommented\{{nameTex}}\{{comment.getString}}\{{hypothesesTex}\n}\{{conclusionTex}}\n" 217 | return s!"\\lottinferencerule\{{nameTex}}\{{hypothesesTex}\n}\{{conclusionTex}}\n" 218 | 219 | let inferenceRulesTex := "\\lottinferencerulesep\n".intercalate inferenceRuleTexs.toList 220 | return s!"\\lottjudgement\{{nameTex}}\{{syntaxTex}}\{\n{inferenceRulesTex}}\n" 221 | 222 | elab_rules : command 223 | | `($jd:Lott.JudgementDecl) => elabJudgementDecls #[jd] 224 | | `(mutual $[$jds:Lott.JudgementDecl]* end) => elabJudgementDecls jds 225 | 226 | @[term_elab zetaReduce] 227 | partial 228 | def elabZetaReduce : TermElab := fun stx expectedType? => do 229 | let `(zeta_reduce% $t) := stx | throwUnsupportedSyntax 230 | let e ← elabTerm t expectedType? 231 | liftMetaM <| Meta.zetaReduce e 232 | 233 | end Lott 234 | -------------------------------------------------------------------------------- /Lott/Data/Range.lean: -------------------------------------------------------------------------------- 1 | import Lott.Data.Nat 2 | 3 | namespace Std.Range 4 | 5 | def toList (r : Range) : List Nat := if r.start < r.stop then 6 | r.start :: toList { r with start := r.start + r.step } 7 | else 8 | [] 9 | termination_by r.stop - r.start 10 | decreasing_by 11 | apply Nat.sub_lt_sub_left _ <| Nat.lt_add_of_pos_right r.step_pos 12 | assumption 13 | 14 | instance : Coe Range (List Nat) := ⟨Std.Range.toList⟩ 15 | 16 | abbrev map (r : Range) (f : Nat → α) : List α := r.toList.map f 17 | 18 | abbrev flatMap (r : Range) (f : Nat → List α) : List α := r.toList.flatMap f 19 | 20 | theorem toList_append (h₁ : l ≤ m) (h₂ : m ≤ n) : [l:m] ++ [m:n] = [l:n].toList := by 21 | rw [toList] 22 | split 23 | · case isTrue h => 24 | apply Eq.symm 25 | rw [toList] 26 | apply Eq.symm 27 | simp only at * 28 | rw [if_pos (Nat.lt_of_lt_of_le h h₂), List.cons_append, toList_append (Nat.succ_le_of_lt h) h₂] 29 | · case isFalse h => 30 | simp only at * 31 | have : l = m := match Nat.lt_trichotomy l m with 32 | | .inl lltm => h lltm |>.elim 33 | | .inr (.inl leqm) => leqm 34 | | .inr (.inr mltn) => Nat.not_lt_of_le h₁ mltn |>.elim 35 | rw [this, List.nil_append] 36 | termination_by m - l 37 | decreasing_by 38 | all_goals simp_wf 39 | apply Nat.sub_succ_lt_self 40 | assumption 41 | 42 | theorem toList_eq_nil_iff : toList [m:n] = [] ↔ n ≤ m where 43 | mp eq := by 44 | rw [toList] at eq 45 | split at eq 46 | · case isTrue h => nomatch eq 47 | · case isFalse h => exact Nat.le_of_not_lt h 48 | mpr le := by rw [toList, if_neg (Nat.not_lt_of_le le)] 49 | 50 | theorem mem_of_mem_toList (h : i ∈ [m:n].toList) : i ∈ [m:n] := by 51 | rw [toList] at h 52 | split at h 53 | · case isTrue mltn => 54 | cases h 55 | · case head => exact ⟨Nat.le_refl _, mltn, Nat.mod_one _⟩ 56 | · case tail h' => 57 | let ⟨msucclei, iltn, _⟩ := mem_of_mem_toList h' 58 | exact ⟨Nat.le_of_succ_le msucclei, iltn, Nat.mod_one _⟩ 59 | · case isFalse => 60 | nomatch h 61 | termination_by n - m 62 | decreasing_by 63 | all_goals simp_arith 64 | apply Nat.sub_succ_lt_self 65 | assumption 66 | 67 | theorem map_eq_of_eq_of_mem {f g : Nat → α} (h : ∀ i ∈ [m:n], f i = g i) 68 | : List.map (fun i => f i) [m:n] = List.map (fun i => g i) [m:n] := 69 | List.map_eq_map_iff.mpr (h · <| mem_of_mem_toList ·) 70 | 71 | theorem map_eq_of_eq_of_mem' {f g : Nat → α} (h : ∀ i ∈ [m:n], f i = g i) 72 | : List.map (fun i => f i) (Coe.coe [m:n]) = List.map (fun i => g i) (Coe.coe [m:n]) := by 73 | dsimp [Coe.coe] 74 | exact map_eq_of_eq_of_mem h 75 | 76 | theorem map_eq_of_eq_of_mem'' {f g : Nat → α} (h : ∀ i ∈ [m:n], f i = g i) 77 | : [m:n].map (fun i => f i) = [m:n].map (fun i => g i) := by exact map_eq_of_eq_of_mem h 78 | 79 | theorem eq_of_mem_of_map_eq {f g : Nat → α} 80 | (h : List.map (fun i => f i) [m:n] = List.map (fun i => g i) [m:n]) : ∀ i ∈ [m:n], f i = g i := by 81 | intro i ⟨mlei, iltn, _⟩ 82 | let mltn := Nat.lt_of_le_of_lt mlei iltn 83 | rw [toList, if_pos mltn] at h 84 | rw [List.map, List.map] at h 85 | simp only at h 86 | by_cases m = i 87 | · case pos h' => 88 | cases h' 89 | exact List.cons_eq_cons.mp h |>.left 90 | · case neg h' => 91 | exact eq_of_mem_of_map_eq (List.cons_eq_cons.mp h |>.right) i 92 | ⟨Nat.succ_le_of_lt <| Nat.lt_of_le_of_ne mlei h', iltn, Nat.mod_one _⟩ 93 | termination_by n - m 94 | decreasing_by 95 | all_goals simp_wf 96 | apply Nat.sub_succ_lt_self 97 | assumption 98 | 99 | theorem eq_of_mem_of_map_eq' {f g : Nat → α} 100 | (h : List.map (fun i => f i) (Coe.coe [m:n]) = List.map (fun i => g i) (Coe.coe [m:n])) 101 | : ∀ i ∈ [m:n], f i = g i := by 102 | dsimp [Coe.coe] at h 103 | exact eq_of_mem_of_map_eq h 104 | 105 | theorem mem_of_mem_map {f : Nat → α} (h : x ∈ List.map (fun i => f i) [m:n]) 106 | : ∃ i ∈ [m:n], x = f i := by 107 | rw [toList] at h 108 | split at h 109 | · case isTrue h' => 110 | rw [List.map] at h 111 | cases h 112 | · case head => exact .intro m ⟨⟨Nat.le_refl _, h', Nat.mod_one _⟩, rfl⟩ 113 | · case tail h'' => 114 | have ⟨i, mem, eq⟩ := mem_of_mem_map (m := m + 1) h'' 115 | exact ⟨i, ⟨Nat.le_of_succ_le mem.lower, mem.upper, Nat.mod_one _⟩, eq⟩ 116 | · case isFalse h' => cases h 117 | termination_by n - m 118 | decreasing_by 119 | all_goals simp_wf 120 | apply Nat.sub_succ_lt_self 121 | assumption 122 | 123 | theorem mem_map_of_mem (h : i ∈ [m:n]) : f i ∈ List.map f [m:n] := by 124 | rw [toList] 125 | split 126 | · case isTrue h' => 127 | simp only 128 | rw [List.map] 129 | by_cases i = m 130 | · case pos h'' => 131 | cases h'' 132 | exact .head _ 133 | · case neg h'' => 134 | exact .tail _ <| mem_map_of_mem ⟨ 135 | Nat.succ_le_of_lt <| Nat.lt_of_le_of_ne h.lower (Ne.symm h''), 136 | h.upper, 137 | Nat.mod_one _ 138 | ⟩ 139 | · case isFalse h' => exact h' (Nat.lt_of_le_of_lt h.lower h.upper) |>.elim 140 | termination_by i - m 141 | decreasing_by 142 | all_goals simp_wf 143 | apply Nat.sub_add_lt_sub _ Nat.one_pos 144 | apply Nat.succ_le_of_lt 145 | apply Nat.lt_of_le_of_ne h.lower 146 | apply Ne.symm 147 | assumption 148 | 149 | theorem map_shift {f : Nat → α} (h : j ≤ m) 150 | : [m - j:n - j].map (fun i => f (i + j)) = [m:n].map fun i => f i := by 151 | rw [map, map, toList] 152 | apply Eq.symm 153 | rw [toList] 154 | apply Eq.symm 155 | simp only 156 | split 157 | · case isTrue h' => 158 | rw [if_pos, List.map, Nat.sub_add_cancel h, ← Nat.sub_add_comm h, ← map, map_shift (m := m + 1), 159 | ← List.map] 160 | · exact Nat.le_succ_of_le h 161 | · have := Nat.sub_pos_of_lt h' 162 | rw [Nat.sub_right_comm, Nat.sub_sub, Nat.sub_add_cancel h] at this 163 | exact Nat.lt_of_sub_pos this 164 | · case isFalse h' => 165 | rw [List.map, if_neg, List.map] 166 | intro h'' 167 | apply h' 168 | have := Nat.sub_pos_of_lt h'' 169 | rw [← Nat.sub_add_cancel h, ← Nat.sub_sub, ← Nat.sub_right_comm] at this 170 | exact Nat.lt_of_sub_pos this 171 | 172 | theorem map_shift' {f : Nat → α} 173 | : [m:n].map (fun i => f i) = [m + j:n + j].map fun i => f (i - j) := by 174 | rw [map, map, toList] 175 | apply Eq.symm 176 | rw [toList] 177 | apply Eq.symm 178 | simp only 179 | split 180 | · case isTrue h' => 181 | rw [if_pos <| Nat.add_lt_add_right h' _, List.map_cons, List.map_cons, Nat.add_sub_self_right, 182 | Nat.add_assoc, Nat.add_comm j 1, ← Nat.add_assoc, ← map, ← map, map_shift'] 183 | · case isFalse h' => 184 | rw [List.map_nil, if_neg <| Nat.not_lt_of_le <| Nat.add_le_add_right (Nat.le_of_not_lt h') _, 185 | List.map_nil] 186 | 187 | theorem map_append {f : Nat → α} (h₁ : l ≤ m) (h₂ : m ≤ n) 188 | : List.map f [l:m] ++ List.map f [m:n] = List.map f [l:n] := by 189 | rw [← List.map_append, toList_append h₁ h₂] 190 | 191 | theorem length_toList : [m:n].toList.length = n - m := by 192 | rw [toList] 193 | split 194 | · case isTrue h => 195 | simp only 196 | rw [List.length_cons, length_toList, Nat.sub_add_eq, Nat.sub_add_cancel] 197 | exact Nat.succ_le_of_lt <| Nat.sub_pos_iff_lt.mpr h 198 | · case isFalse h => rw [List.length_nil, Nat.sub_eq_zero_of_le <| Nat.le_of_not_lt h] 199 | termination_by n - m 200 | decreasing_by 201 | all_goals simp_arith 202 | apply Nat.sub_succ_lt_self 203 | assumption 204 | 205 | theorem getElem_toList (mlt : m < n - l) 206 | : [l:n].toList[m]'(by rw [length_toList]; exact mlt) = l + m := by 207 | rw [List.getElem_eq_iff] 208 | rw [toList] 209 | split 210 | · case isTrue h => 211 | simp only 212 | rw [List.getElem?_cons] 213 | split 214 | · case isTrue h => 215 | cases h 216 | rfl 217 | · case isFalse h => 218 | have := getElem_toList (l := l + 1) (m := m - 1) (n := n) <| by 219 | apply Nat.lt_sub_of_add_lt 220 | rw [Nat.add_comm l, ← Nat.add_assoc, Nat.sub_add_cancel <| Nat.pos_of_ne_zero h] 221 | exact Nat.add_lt_of_lt_sub mlt 222 | rw [List.getElem_eq_iff] at this 223 | rw [Nat.add_assoc l 1 (m - 1), Nat.add_comm 1 (m - 1), 224 | Nat.sub_add_cancel <| Nat.pos_of_ne_zero h] at this 225 | exact this 226 | · case isFalse h => 227 | simp at h 228 | have := Nat.add_lt_of_lt_sub mlt 229 | rw [Nat.add_comm] at this 230 | nomatch Nat.not_le.mpr (Nat.lt_of_add_right_lt this) h 231 | 232 | theorem map_get!_eq [Inhabited α] {as : List α} : [:as.length].map as.get! = as := by 233 | match as with 234 | | [] => 235 | rw [List.length_nil, map, toList, if_neg (Nat.not_lt_of_le (Nat.le_refl _)), List.map_nil] 236 | | a :: as' => 237 | rw [List.length_cons, map, toList, if_pos (Nat.succ_pos _), List.map_cons, List.get!_cons_zero, 238 | ← map, ← map_shift (Nat.le_add_left ..), Nat.add_sub_cancel, Nat.add_sub_cancel, 239 | map_eq_of_eq_of_mem'' fun _ _ => List.get!_cons_succ .., map_get!_eq] 240 | 241 | theorem count_toList_le_one : [m:n].toList.count l ≤ 1 := by 242 | rw [toList] 243 | split 244 | · case isTrue h => 245 | rw [List.count_cons] 246 | simp only 247 | split 248 | · case isTrue h' => 249 | cases Nat.eq_of_beq_eq_true' h' 250 | rw [List.count_eq_zero_of_not_mem] 251 | · exact Nat.le_refl _ 252 | · intro lmem 253 | nomatch Nat.ne_of_lt <| Nat.lt_of_succ_le <| And.left <| mem_of_mem_toList lmem 254 | · case isFalse h' => 255 | rw [Nat.add_zero] 256 | exact count_toList_le_one 257 | · case isFalse h => 258 | rw [List.count_nil] 259 | exact Nat.le_of_lt Nat.one_pos 260 | termination_by n - m 261 | decreasing_by 262 | all_goals simp_arith 263 | apply Nat.sub_succ_lt_self 264 | assumption 265 | 266 | theorem get!_map [Inhabited α] {f : Nat → α} (iltnsubm : i < n - m) 267 | : ([m:n].map f).get! i = f (i + m) := by match i with 268 | | 0 => 269 | rw [map, toList, if_pos (Nat.lt_of_sub_pos iltnsubm), List.map_cons, List.get!_cons_zero, 270 | Nat.zero_add] 271 | | i' + 1 => 272 | let mltn := Nat.lt_of_sub_pos (Nat.lt_of_le_of_lt (Nat.zero_le _) iltnsubm) 273 | rw [map, toList, if_pos mltn, List.map_cons, List.get!_cons_succ, ← map, 274 | ← map_shift (j := 1) (Nat.succ_le_of_lt (Nat.add_one_pos _)), get!_map, Nat.add_sub_cancel, 275 | Nat.add_assoc, Nat.add_comm m, ← Nat.add_assoc] 276 | rw [Nat.add_sub_cancel, Nat.sub_right_comm] 277 | apply Nat.lt_sub_of_add_lt 278 | exact iltnsubm 279 | 280 | theorem skolem [Inhabited α] {p : Nat → α → Prop} 281 | : (∀ i ∈ [:n], ∃ a, p i a) → ∃ f : Nat → α, ∀ i ∈ [:n], p i (f i) := by 282 | intro h 283 | induction n with 284 | | zero => exact ⟨fun _ => default, nofun⟩ 285 | | succ n ih => 286 | let ⟨f', h'⟩ := ih fun i mem => h i ⟨mem.lower, Nat.lt_add_right _ mem.upper, Nat.mod_one _⟩ 287 | let ⟨a, h''⟩ := h n ⟨Nat.zero_le _, Nat.lt_succ_self _, Nat.mod_one _⟩ 288 | let f i := if i = n then a else f' i 289 | apply Exists.intro f 290 | intro i mem 291 | dsimp only [f] 292 | split 293 | · case isTrue h''' => 294 | cases h''' 295 | exact h'' 296 | · case isFalse h''' => 297 | exact h' i ⟨mem.lower, Nat.lt_of_le_of_ne (Nat.le_of_lt_succ mem.upper) h''', Nat.mod_one _⟩ 298 | 299 | theorem mem_zip_map {f : Nat → α} {g : Nat → β} 300 | : x ∈ ([m:n].map f).zip ([m:n].map g) → ∃ i ∈ [m:n], x = (f i, g i) := by 301 | intro mem 302 | rw [map, map, toList] at mem 303 | split at mem 304 | case isFalse h => nomatch mem 305 | case isTrue h => 306 | rw [List.map_cons, List.map_cons, List.zip_cons_cons] at mem 307 | cases mem with 308 | | head => exact ⟨m, ⟨Nat.le.refl, h, Nat.mod_one _⟩, rfl⟩ 309 | | tail _ mem' => 310 | let ⟨i, imem, eq⟩ := mem_zip_map mem' 311 | simp at * 312 | exact ⟨i, ⟨Nat.le_trans Nat.le.refl.step imem.lower, imem.upper, Nat.mod_one _⟩, eq⟩ 313 | termination_by n - m 314 | decreasing_by 315 | all_goals simp_arith 316 | apply Nat.sub_succ_lt_self 317 | assumption 318 | 319 | theorem map_eq_cons_of_lt (mltn : m < n) : [m:n].map f = f m :: [m+1:n].map f := by 320 | rw [map, toList, if_pos mltn, List.map_cons, ← map] 321 | 322 | theorem map_same_eq_nil : [n:n].map f = [] := by 323 | rw [map, toList, if_neg <| Nat.not_lt_of_le Nat.le.refl, List.map_nil] 324 | 325 | theorem map_eq_snoc_of_lt (mltn : m < n) : [m:n].map f = [m:n - 1].map f ++ [f (n - 1)] := by 326 | let npos := Nat.lt_of_le_of_lt (Nat.zero_le _) mltn 327 | rw [map, ← map_append (l := m) (m := n - 1) (n := n) (Nat.le_sub_one_of_lt mltn) (Nat.sub_le _ _), 328 | ← map, ← map, map_eq_cons_of_lt <| Nat.sub_lt npos Nat.one_pos, Nat.sub_add_cancel npos, 329 | map_same_eq_nil] 330 | 331 | end Std.Range 332 | -------------------------------------------------------------------------------- /LottExamples/STLC/Proofs.lean: -------------------------------------------------------------------------------- 1 | import Aesop 2 | import Lott.Data.List 3 | import LottExamples.STLC.Semantics 4 | 5 | namespace LottExamples.STLC 6 | 7 | namespace Term 8 | 9 | @[simp] 10 | theorem Var_open_sizeOf : sizeOf (Var_open e x n) = sizeOf e := by 11 | induction e generalizing n <;> aesop (add simp Var_open) 12 | 13 | theorem Var_open_comm (e : Term) 14 | : m ≠ n → (e.Var_open x m).Var_open x' n = (e.Var_open x' n).Var_open x m := by 15 | induction e generalizing m n <;> aesop (add simp Var_open) 16 | 17 | namespace VarLocallyClosed 18 | 19 | theorem weakening (elc : VarLocallyClosed e m) : m ≤ n → e.VarLocallyClosed n := by 20 | induction elc generalizing n 21 | all_goals aesop (add 20% constructors VarLocallyClosed, 20% Nat.lt_of_lt_of_le) 22 | 23 | theorem Var_open_drop : m < n → (Var_open e x m).VarLocallyClosed n → e.VarLocallyClosed n := by 24 | induction e generalizing m n <;> aesop 25 | (add simp Var_open, safe cases VarLocallyClosed, safe constructors VarLocallyClosed) 26 | 27 | theorem Var_open_id : VarLocallyClosed e n → e.Var_open x n = e := by 28 | induction e generalizing n <;> aesop (add simp Var_open, safe cases VarLocallyClosed) 29 | 30 | theorem Var_open_Term_open_comm (e'lc : VarLocallyClosed e') 31 | : m ≠ n → (Term_open e e' m).Var_open x n = (e.Var_open x n).Term_open e' m := by 32 | induction e generalizing m n <;> aesop 33 | (add simp [Var_open, Term_open], 50% apply Nat.zero_le, 20% apply [Var_open_id, weakening]) 34 | 35 | end VarLocallyClosed 36 | 37 | theorem InFreeVars.Var_open_drop : InFreeVars x (e.Var_open x' n) → x ≠ x' → [[x ∈ fv(e)]] := by 38 | induction e generalizing n <;> all_goals aesop (add simp [InFreeVars, Var_open, freeVars]) 39 | 40 | namespace NotInFreeVars 41 | 42 | theorem lam : [[x ∉ fv(λ x. e)]] → [[x ∉ fv(e)]] := (· ·) 43 | 44 | theorem app₀ : [[x ∉ fv(e₀ e₁)]] → [[x ∉ fv(e₀)]] := (· <| List.mem_append.mpr <| .inl ·) 45 | 46 | theorem app₁ : [[x ∉ fv(e₀ e₁)]] → [[x ∉ fv(e₁)]] := (· <| List.mem_append.mpr <| .inr ·) 47 | 48 | theorem Var_open_intro : [[x ∉ fv(e)]] → x ≠ x' → [[x ∉ fv(e^x')]] := 49 | fun xninfve xnex' xinfveop => xninfve <| xinfveop.Var_open_drop xnex' 50 | 51 | end NotInFreeVars 52 | 53 | end Term 54 | 55 | namespace Environment 56 | 57 | theorem VarNotIn.ext : [[x ∉ Γ, x' : τ]] ↔ x ≠ x' ∧ [[x ∉ Γ]] where 58 | mp xninΓx' := ⟨ 59 | fun | .refl .. => xninΓx' _ .head, 60 | fun _ xinΓ => Classical.byCases 61 | fun | .refl .. => xninΓx' _ .head 62 | fun xnex' => xninΓx' _ <| xinΓ.ext xnex' 63 | ⟩ 64 | mpr 65 | | ⟨xnex', _⟩, _, .head => nomatch xnex' 66 | | ⟨_, xninΓ⟩, _, .ext xinΓ _ => xninΓ _ xinΓ 67 | 68 | namespace VarIn 69 | 70 | theorem append_elim : [[x : τ ∈ Γ₀, Γ₁]] → [[x : τ ∈ Γ₀]] ∧ [[x ∉ Γ₁]] ∨ [[x : τ ∈ Γ₁]] := 71 | fun xinΓ₀Γ₁ => match Γ₁ with 72 | | .empty => .inl ⟨xinΓ₀Γ₁, nofun⟩ 73 | | .ext .. => 74 | match xinΓ₀Γ₁ with 75 | | .head => .inr head 76 | | .ext xinΓ₀Γ₁' xnex' => match xinΓ₀Γ₁'.append_elim with 77 | | .inl ⟨xinΓ₀, xninΓ₁'⟩ => .inl ⟨xinΓ₀, VarNotIn.ext.mpr ⟨xnex', xninΓ₁'⟩⟩ 78 | | .inr xinΓ₁' => .inr <| xinΓ₁'.ext xnex' 79 | 80 | theorem append_inl : [[x : τ ∈ Γ₀]] → [[x ∉ Γ₁]] → [[x : τ ∈ Γ₀, Γ₁]] := 81 | fun xinΓ₀ xninΓ₁ => match Γ₁ with 82 | | .empty => xinΓ₀ 83 | | .ext .. => 84 | let ⟨xnex', xninΓ₁'⟩ := VarNotIn.ext.mp xninΓ₁ 85 | xinΓ₀.append_inl xninΓ₁' |>.ext xnex' 86 | 87 | theorem append_inr : [[x : τ ∈ Γ₁]] → [[x : τ ∈ Γ₀, Γ₁]] 88 | | head => head 89 | | ext xinΓ₁' xnex' => xinΓ₁'.append_inr.ext xnex' 90 | 91 | theorem VarInDom_of : [[x : τ ∈ Γ]] → [[x ∈ dom(Γ)]] 92 | | .head => .head _ 93 | | .ext xinΓ' ne => .tail _ xinΓ'.VarInDom_of 94 | 95 | end VarIn 96 | 97 | theorem VarNotIn.append : [[x ∉ Γ₀, Γ₁]] ↔ [[x ∉ Γ₀]] ∧ [[x ∉ Γ₁]] where 98 | mp xninΓ₀Γ₁ := ⟨ 99 | fun _ xinΓ₀ => Classical.byCases (p := ∃ τ, [[x : τ ∈ Γ₁]]) 100 | (fun ⟨_, xinΓ₁⟩ => xninΓ₀Γ₁ _ xinΓ₁.append_inr) 101 | fun xninΓ₁ => xninΓ₀Γ₁ _ <| xinΓ₀.append_inl <| not_exists.mp xninΓ₁, 102 | fun _ => (xninΓ₀Γ₁ _ ·.append_inr) 103 | ⟩ 104 | mpr | ⟨xninΓ₀, xninΓ₁⟩, _, xinΓ₀Γ₁ => match xinΓ₀Γ₁.append_elim with 105 | | .inl ⟨xinΓ₀, _⟩ => xninΓ₀ _ xinΓ₀ 106 | | .inr xinΓ₁ => xninΓ₁ _ xinΓ₁ 107 | 108 | namespace VarInDom 109 | 110 | theorem insert : [[x ∈ dom(Γ₀, Γ₁)]] → [[x ∈ dom(Γ₀, x' : τ, Γ₁)]] := 111 | fun xindomΓ₀Γ₁ => 112 | match Γ₁ with 113 | | .empty => .tail _ xindomΓ₀Γ₁ 114 | | .ext .. => match xindomΓ₀Γ₁ with 115 | | .head _ => .head _ 116 | | .tail _ xindomΓ₀Γ₁' => .tail _ <| insert xindomΓ₀Γ₁' 117 | 118 | theorem of_VarIn : [[x : τ ∈ Γ]] → [[x ∈ dom(Γ)]] 119 | | .head => .head _ 120 | | .ext xinΓ' _ => .tail _ <| of_VarIn xinΓ' 121 | 122 | theorem append_elim : [[x ∈ dom(Γ₀, Γ₁)]] → [[x ∈ dom(Γ₀)]] ∨ [[x ∈ dom(Γ₁)]] := fun xindomΓ₀Γ₁ => 123 | match Γ₁ with 124 | | .empty => .inl xindomΓ₀Γ₁ 125 | | .ext .. => match xindomΓ₀Γ₁ with 126 | | .head _ => .inr <| .head _ 127 | | .tail _ xindomΓ₀Γ₁' => match append_elim xindomΓ₀Γ₁' with 128 | | .inl xindomΓ₀ => .inl xindomΓ₀ 129 | | .inr xindomΓ₁' => .inr <| .tail _ <| xindomΓ₁' 130 | 131 | theorem append_inl : [[x ∈ dom(Γ₀)]] → [[x ∈ dom(Γ₀, Γ₁)]] := fun xindomΓ₀ => match Γ₁ with 132 | | .empty => xindomΓ₀ 133 | | .ext .. => .tail _ xindomΓ₀.append_inl 134 | 135 | theorem append_inr : [[x ∈ dom(Γ₁)]] → [[x ∈ dom(Γ₀, Γ₁)]] := fun xindomΓ₁ => 136 | let .ext .. := Γ₁ 137 | match xindomΓ₁ with 138 | | .head _ => .head _ 139 | | .tail _ xinΓ₁' => .tail _ <| append_inr xinΓ₁' 140 | 141 | end VarInDom 142 | 143 | namespace VarNotInDom 144 | 145 | theorem drop : [[x ∉ dom(Γ₀, x' : τ, Γ₁)]] → [[x ∉ dom(Γ₀, Γ₁)]] := 146 | (· ·.insert) 147 | 148 | theorem append : [[x ∉ dom(Γ₀, Γ₁)]] ↔ [[x ∉ dom(Γ₀)]] ∧ [[x ∉ dom(Γ₁)]] where 149 | mp xnindomΓ₀Γ₁ := ⟨(xnindomΓ₀Γ₁ ·.append_inl), (xnindomΓ₀Γ₁ ·.append_inr)⟩ 150 | mpr | ⟨xnindomΓ₀, xnindomΓ₁⟩, xindomΓ₀Γ₁ => match xindomΓ₀Γ₁.append_elim with 151 | | .inl xindomΓ₀ => xnindomΓ₀ xindomΓ₀ 152 | | .inr xindomΓ₁ => xnindomΓ₁ xindomΓ₁ 153 | 154 | theorem ext : [[x ∉ dom(Γ, x' : τ)]] ↔ x ≠ x' ∧ [[x ∉ dom(Γ)]] where 155 | mp xnindomΓx' := ⟨ 156 | fun | .refl .. => xnindomΓx' <| .head _, 157 | (xnindomΓx' <| .tail _ ·) 158 | ⟩ 159 | mpr | ⟨xnex', xnindomΓ⟩, xindomΓx' => match xindomΓx' with 160 | | .head _ => nomatch xnex' 161 | | .tail _ xindomΓ => xnindomΓ xindomΓ 162 | 163 | theorem exchange : [[x ∉ dom(Γ₀, x' : τ, Γ₁, Γ₂)]] → [[x ∉ dom(Γ₀, Γ₁, x' : τ, Γ₂)]] := 164 | fun xnindomΓ₀x'Γ₁Γ₂ => 165 | let ⟨xnindomΓ₀x', xnindomΓ₁Γ₂⟩ := append.mp xnindomΓ₀x'Γ₁Γ₂ 166 | let ⟨xnex', xnindomΓ₀⟩ := ext.mp xnindomΓ₀x' 167 | let ⟨xnindomΓ₁, xnindomΓ₂⟩ := append.mp xnindomΓ₁Γ₂ 168 | append.mpr ⟨xnindomΓ₀, append.mpr ⟨ext.mpr ⟨xnex', xnindomΓ₁⟩, xnindomΓ₂⟩⟩ 169 | 170 | theorem not_VarIn : [[x ∉ dom(Γ)]] → [[x ∉ Γ]] := (· <| VarIn.VarInDom_of (τ := ·) ·) 171 | 172 | end VarNotInDom 173 | 174 | namespace WellFormedness 175 | 176 | theorem insert : [[⊢ Γ₀, Γ₁]] → [[x ∉ dom(Γ₀, Γ₁)]] → [[⊢ Γ₀, x : τ, Γ₁]] := 177 | fun Γ₀Γ₁wf xnindomΓ₀Γ₁ => by match Γ₁ with 178 | | [[ε]] => exact Γ₀Γ₁wf.ext xnindomΓ₀Γ₁ 179 | | [[Γ₁', x' : τ']] => 180 | let .ext Γ₀Γ₁'wf x'nindomΓ₀Γ₁' := Γ₀Γ₁wf 181 | let ⟨x'nindomΓ₀, x'nindomΓ₁'⟩ := VarNotInDom.append.mp x'nindomΓ₀Γ₁' 182 | let ⟨_, xnindomΓ₁⟩ := VarNotInDom.append.mp xnindomΓ₀Γ₁ 183 | let x'nindom : [[x' ∉ dom(Γ₀, x : τ, Γ₁')]] := VarNotInDom.append.mpr 184 | ⟨VarNotInDom.ext.mpr ⟨VarNotInDom.ext.mp xnindomΓ₁ |>.left.symm, x'nindomΓ₀⟩, x'nindomΓ₁'⟩ 185 | exact Γ₀Γ₁'wf.insert (VarNotInDom.ext.mp xnindomΓ₀Γ₁).right |>.ext x'nindom 186 | 187 | theorem drop : [[⊢ Γ₀, x : τ, Γ₁]] → [[⊢ Γ₀, Γ₁]] := 188 | fun Γ₀xΓ₁wf => match Γ₁ with 189 | | .empty => 190 | let .ext Γ₀wf _ := Γ₀xΓ₁wf 191 | Γ₀wf 192 | | .ext .. => 193 | let .ext Γ₀xΓ₁'wf xnindomΓ₀xΓ₁' := Γ₀xΓ₁wf 194 | Γ₀xΓ₁'wf.drop.ext xnindomΓ₀xΓ₁'.drop 195 | 196 | theorem exchange : [[⊢ Γ₀, x : τ, Γ₁, Γ₂]] → [[⊢ Γ₀, Γ₁, x : τ, Γ₂]] := fun Γ₀xΓ₁Γ₂wf => 197 | match Γ₂ with 198 | | [[ε]] => by induction Γ₁ with 199 | | empty => exact Γ₀xΓ₁Γ₂wf 200 | | ext Γ₁' x' τ' ih => 201 | simp [append] at Γ₀xΓ₁Γ₂wf ih ⊢ 202 | let .ext Γ₀xΓ₁'wf x'nindomΓ₀xΓ₁' := Γ₀xΓ₁Γ₂wf 203 | let .ext Γ₀Γ₁'wf xnindomΓ₀Γ₁ := ih Γ₀xΓ₁'wf 204 | let ⟨x'nindomΓ₀x, _⟩ := VarNotInDom.append.mp x'nindomΓ₀xΓ₁' 205 | let ⟨xnex', _⟩ := VarNotInDom.ext.mp x'nindomΓ₀x 206 | exact Γ₀Γ₁'wf.ext x'nindomΓ₀xΓ₁'.drop |>.ext <| VarNotInDom.ext.mpr ⟨xnex'.symm, xnindomΓ₀Γ₁⟩ 207 | | [[Γ₂', x' : τ']] => 208 | let .ext Γ₀xΓ₁Γ₂'wf x'ninΓ₀xΓ₁Γ₂' := Γ₀xΓ₁Γ₂wf 209 | Γ₀xΓ₁Γ₂'wf.exchange.ext x'ninΓ₀xΓ₁Γ₂'.exchange 210 | 211 | theorem append_elim : [[⊢ Γ₀, Γ₁]] → [[⊢ Γ₀]] ∧ [[⊢ Γ₁]] := fun Γ₀Γ₁wf => 212 | match Γ₁ with 213 | | .empty => ⟨Γ₀Γ₁wf, .empty⟩ 214 | | .ext .. => 215 | let .ext Γ₀Γ₁'wf xninΓ₀Γ₁' := Γ₀Γ₁wf 216 | let ⟨Γ₀wf, Γ₁'wf⟩ := Γ₀Γ₁'wf.append_elim 217 | ⟨Γ₀wf, .ext Γ₁'wf <| VarNotInDom.append.mp xninΓ₀Γ₁' |>.right⟩ 218 | 219 | end WellFormedness 220 | 221 | theorem VarIn.exchange 222 | : [[x : τ ∈ Γ₀, x' : τ', Γ₁, Γ₂]] → [[⊢ Γ₀, x' : τ', Γ₁, Γ₂]] → [[x : τ ∈ Γ₀, Γ₁, x' : τ', Γ₂]] := 223 | fun xinΓ₀x'Γ₁Γ₂ Γ₀x'Γ₁Γ₂wf => 224 | match xinΓ₀x'Γ₁Γ₂.append_elim with 225 | | .inl ⟨xinΓ₀x', xninΓ₁Γ₂⟩ => 226 | let ⟨xninΓ₁, xninΓ₂⟩ := VarNotIn.append.mp xninΓ₁Γ₂ 227 | match xinΓ₀x' with 228 | | .head => VarIn.head.append_inl xninΓ₂ |>.append_inr 229 | | .ext xinΓ₀ xnex' => 230 | xinΓ₀.append_inl <| VarNotIn.append.mpr ⟨VarNotIn.ext.mpr ⟨xnex', xninΓ₁⟩, xninΓ₂⟩ 231 | | .inr xinΓ₁Γ₂ => match xinΓ₁Γ₂.append_elim with 232 | | .inl ⟨xinΓ₁, xninΓ₂⟩ => 233 | let f xeqx' := by 234 | subst xeqx' 235 | rw [Environment.append_assoc] at Γ₀x'Γ₁Γ₂wf 236 | let .ext _ xninΓ₀Γ₁ := Γ₀x'Γ₁Γ₂wf.append_elim.left.exchange (Γ₂ := .empty) 237 | let xninΓ₁ := VarNotInDom.append.mp xninΓ₀Γ₁ |>.right 238 | nomatch xninΓ₁.not_VarIn _ xinΓ₁ 239 | Classical.byCases (p := x = x') f (xinΓ₁.ext · |>.append_inl xninΓ₂) |>.append_inr 240 | | .inr xinΓ₂ => xinΓ₂.append_inr.append_inr 241 | 242 | theorem VarNotIn.of_VarIn_of_WellFormedness 243 | : [[x : τ ∈ Γ₀]] → [[⊢ Γ₀, Γ₁]] → [[x ∉ Γ₁]] := fun xinΓ₀ Γ₀Γ₁wf => match Γ₁ with 244 | | .empty => nofun 245 | | .ext .. => 246 | let .ext Γ₀Γ₁'wf x'ninΓ₀Γ₁ := Γ₀Γ₁wf 247 | ext.mpr ⟨ 248 | fun | .refl .. => (VarNotInDom.append.mp x'ninΓ₀Γ₁).left <| VarInDom.of_VarIn xinΓ₀, 249 | of_VarIn_of_WellFormedness xinΓ₀ Γ₀Γ₁'wf 250 | ⟩ 251 | 252 | end Environment 253 | 254 | namespace Typing 255 | 256 | theorem toVarLocallyClosed : [[Γ ⊢ e : τ]] → e.VarLocallyClosed 257 | | var .. => .var_free 258 | | lam e'ty (I := I) => 259 | let ⟨x, xnin⟩ := I.exists_fresh 260 | let e'ty := e'ty x xnin 261 | .lam <| e'ty.toVarLocallyClosed.weakening (Nat.le_succ 0) |>.Var_open_drop <| Nat.zero_lt_succ _ 262 | | app e₀ty e₁ty => .app e₀ty.toVarLocallyClosed e₁ty.toVarLocallyClosed 263 | | unit => .unit 264 | 265 | theorem exchange : [[Γ₀, x : τ, Γ₁, Γ₂ ⊢ e : τ']] → [[Γ₀, Γ₁, x : τ, Γ₂ ⊢ e : τ']] 266 | | var Γ₀xΓ₁Γ₂wf x'inΓ₀xΓ₁Γ₂ => var Γ₀xΓ₁Γ₂wf.exchange <| x'inΓ₀xΓ₁Γ₂.exchange Γ₀xΓ₁Γ₂wf 267 | | lam e'ty => lam fun x' x'nin => let e'ty := e'ty x' x'nin; e'ty.exchange (Γ₂ := Γ₂.ext x' _) 268 | | app e₀ty e₁ty => app e₀ty.exchange e₁ty.exchange 269 | | unit => unit 270 | 271 | theorem weakening : [[Γ₀ ⊢ e : τ]] → [[⊢ Γ₀, Γ₁]] → [[Γ₀, Γ₁ ⊢ e : τ]] 272 | | var _ xinΓ₀, Γ₀Γ₁wf => 273 | var Γ₀Γ₁wf <| xinΓ₀.append_inl <| Environment.VarNotIn.of_VarIn_of_WellFormedness xinΓ₀ Γ₀Γ₁wf 274 | | lam e'ty (I := I), Γ₀Γ₁wf => 275 | lam (I := (Γ₀.append Γ₁).domain ++ I) fun x xnin => by 276 | let ⟨xnindomΓ₀Γ₁, xninI⟩ := List.not_mem_append'.mp xnin 277 | let e'ty := e'ty x xninI 278 | have := e'ty.weakening (Γ₀Γ₁wf.insert xnindomΓ₀Γ₁) 279 | exact this.exchange (Γ₂ := .empty) 280 | | app e₀ty e₁ty, Γ₀Γ₁wf => app (e₀ty.weakening Γ₀Γ₁wf) (e₁ty.weakening Γ₀Γ₁wf) 281 | | unit, _ => unit 282 | 283 | theorem opening 284 | (e₁ty : Typing ((Γ₀.ext x τ₀).append Γ₁) (e₁.Var_open x n) τ₁) (e₀ty : [[Γ₀ ⊢ e₀ : τ₀]]) 285 | (xninΓ₁ : [[x ∉ Γ₁]]) (xninfve₁ : [[x ∉ fv(e₁)]]) 286 | : Typing (Γ₀.append Γ₁) (e₁.Term_open e₀ n) τ₁ := by 287 | match e₁ with 288 | | .var (.free x') => 289 | rw [Term.Var_open, if_neg nofun] at e₁ty 290 | let .var Γ₀xΓ₁wf x'inΓ₀xΓ₁ := e₁ty 291 | match x'inΓ₀xΓ₁.append_elim with 292 | | .inl ⟨.head, x'ninΓ₁⟩ => nomatch List.not_mem_singleton.mp xninfve₁ 293 | | .inl ⟨.ext x'inΓ₀ _, x'ninΓ₁⟩ => exact .var Γ₀xΓ₁wf.drop <| x'inΓ₀.append_inl x'ninΓ₁ 294 | | .inr x'inΓ₁ => exact .var Γ₀xΓ₁wf.drop x'inΓ₁.append_inr 295 | | .var (.bound _) => 296 | rw [Term.Var_open] at e₁ty 297 | split at e₁ty 298 | · case isTrue h => 299 | cases h 300 | rw [Term.Term_open, if_pos rfl] 301 | let .var Γ₀xΓ₁wf xinΓ₀xΓ₁ := e₁ty 302 | match xinΓ₀xΓ₁.append_elim with 303 | | .inl ⟨.head, _⟩ => exact e₀ty.weakening Γ₀xΓ₁wf.drop 304 | | .inr xinΓ₁ => exact xninΓ₁ _ xinΓ₁ |>.elim 305 | · case isFalse h => nomatch e₁ty 306 | | [[λ x. e₁']] => 307 | rw [Term.Term_open] 308 | let .lam e₁'ty (τ₀ := τ₀') (I := I) := e₁ty 309 | exact .lam (I := x :: I) fun x' x'nin => by 310 | let xnex' := List.ne_of_not_mem_cons x'nin 311 | let e₁'ty := e₁'ty x' <| List.not_mem_of_not_mem_cons x'nin 312 | have : ((Γ₀.ext x τ₀).append Γ₁).ext x' τ₀' = (Γ₀.ext x τ₀).append (Γ₁.ext x' τ₀') := rfl 313 | rw [this, e₁'.Var_open_comm <| Nat.succ_ne_zero _] at e₁'ty 314 | rw [e₀ty.toVarLocallyClosed.Var_open_Term_open_comm <| Nat.succ_ne_zero _] 315 | let xninΓ₁x' : [[x ∉ Γ₁, x' : τ₀']] := Environment.VarNotIn.ext.mpr ⟨xnex'.symm, xninΓ₁⟩ 316 | exact e₁'ty.opening e₀ty xninΓ₁x' <| xninfve₁.lam.Var_open_intro xnex'.symm 317 | | [[e₁₀ e₁₁]] => 318 | let .app e₁₀ty e₁₁ty := e₁ty 319 | exact .app (e₁₀ty.opening e₀ty xninΓ₁ xninfve₁.app₀) (e₁₁ty.opening e₀ty xninΓ₁ xninfve₁.app₁) 320 | | [[()]] => 321 | let .unit := e₁ty 322 | exact unit 323 | 324 | end Typing 325 | 326 | namespace Reduction 327 | 328 | theorem preservation (ty : [[Γ ⊢ e : τ]]) (re : [[e ↦ e']]) : [[Γ ⊢ e' : τ]] := match re, ty with 329 | | appl e₀ree₀', .app e₀ty e₁ty => .app (e₀ree₀'.preservation e₀ty) e₁ty 330 | | appr e₁ree₁', .app e₀ty e₁ty => .app e₀ty <| e₁ree₁'.preservation e₁ty 331 | | lamApp, .app e₀ty vty => 332 | let .lam e₀'ty (e := e₀') (I := I) := e₀ty 333 | let ⟨x, xnin⟩ := e₀'.freeVars ++ I |>.exists_fresh 334 | let ⟨xninfve₀', xninI⟩ := List.not_mem_append'.mp xnin 335 | e₀'ty x xninI |>.opening (Γ₁ := .empty) vty nofun xninfve₀' 336 | 337 | theorem progress (ty : [[ε ⊢ e : τ]]) : e.IsValue ∨ ∃ e', [[e ↦ e']] := match e, ty with 338 | | [[λ x. e₀]], _ => .inl .lam 339 | | [[e₀ e₁]], .app e₀ty e₁ty => match progress e₀ty with 340 | | .inl e₀IsValue => match progress e₁ty with 341 | | .inl e₁IsValue => 342 | let [[λ x. e₀']] := e₀ 343 | let v₁ : Value := ⟨e₁, e₁IsValue⟩ 344 | .inr <| .intro _ <| .lamApp (v := v₁) 345 | | .inr ⟨_, e₁ree₁'⟩ => .inr <| .intro _ <| .appr e₁ree₁' (v := ⟨_, e₀IsValue⟩) 346 | | .inr ⟨_, e₀ree₀'⟩ => .inr <| .intro _ <| .appl e₀ree₀' 347 | | [[()]], .unit => .inl .unit 348 | 349 | end Reduction 350 | 351 | end LottExamples.STLC 352 | -------------------------------------------------------------------------------- /Lott/Elab/Basic.lean: -------------------------------------------------------------------------------- 1 | import Lott.Data.Char 2 | import Lott.Data.List 3 | import Lott.Data.Name 4 | import Lott.Data.Option 5 | import Lott.Data.String 6 | import Lott.Elab.Options 7 | import Lott.Environment 8 | import Lott.Parser 9 | import Lott.IR 10 | 11 | -- TODO: Remove unnecessary qualifications from names. 12 | -- TODO: Delab to embeddings when possible. 13 | -- TODO: Make hover in non-terminal work right. 14 | -- TODO: Test stuff without locally nameless. 15 | -- TODO: Centralize validation of bind/id usage in nonterminals. 16 | 17 | namespace Lott 18 | 19 | open IO.FS 20 | open Lean 21 | open Lean.Elab 22 | open Lean.Elab.Command 23 | open Lean.Elab.Term 24 | open Lean.Syntax 25 | open Lean.Parser.Command 26 | open Lean.Parser.Term 27 | open Lott.IR 28 | open System 29 | 30 | /- Utility stuff. -/ 31 | 32 | def writeMakeDeps (filePath : FilePath) (extraDeps : List FilePath := []) : CommandElabM Unit := do 33 | if !(← getOptions).get lott.tex.output.makeDeps.name lott.tex.output.makeDeps.defValue then 34 | return 35 | 36 | let sp ← liftIO <| initSrcSearchPath 37 | let depPaths ← (← getEnv).allImportedModuleNames.mapM (findLean sp ·) 38 | let deps := " ".intercalate <| (depPaths.toList ++ extraDeps).map 39 | (·.toString.dropPrefixes "./" |>.toString) 40 | writeFile (filePath.addExtension "d") 41 | s!"{filePath.toString.dropPrefixes "./" |>.toString}: {deps}\n" 42 | 43 | def writeTexOutput (name : Name) (mkTex : CommandElabM String) : CommandElabM Unit := do 44 | let opts ← getOptions 45 | let some (outputDir : String) := opts.get? lott.tex.output.dir.name | return 46 | let outputDir ← if FilePath.isAbsolute outputDir then 47 | pure (outputDir : FilePath) 48 | else if opts.get lott.tex.output.sourceRelative.name lott.tex.output.sourceRelative.defValue then 49 | let some parent := FilePath.parent <| ← getFileName 50 | | throwError "failed to resolve parent of current file" 51 | pure <| parent / outputDir 52 | else 53 | pure outputDir 54 | let outputName := outputDir / name.toStringWithSep "/" false |>.addExtension "tex" 55 | let some parent := outputName.parent | throwError "failed to resolve parent of output file name" 56 | createDirAll parent 57 | 58 | writeMakeDeps outputName 59 | writeFile outputName <| ← mkTex 60 | 61 | abbrev TexElab := (profile : Name) → (ref : Syntax) → Syntax → TermElabM String 62 | 63 | private unsafe 64 | def mkLottTexElabAttributeUnsafe (ref : Name) : IO (KeyedDeclsAttribute TexElab) := do 65 | mkElabAttribute TexElab .anonymous `lott_tex_elab `Lott ``Lott.TexElab "lott" ref 66 | 67 | @[implemented_by mkLottTexElabAttributeUnsafe] 68 | opaque mkLottTexElabAttribute (ref : Name) : IO (KeyedDeclsAttribute TexElab) 69 | 70 | initialize lottTexElabAttribute : KeyedDeclsAttribute TexElab ← mkLottTexElabAttribute decl_name% 71 | 72 | def texElabIdx : Syntax → TermElabM String 73 | | .ident _ _ val _ => pure <| val.toString false |>.texEscape 74 | | .node _ `num #[.atom _ raw] => pure raw.texEscape 75 | | _ => throwUnsupportedSyntax 76 | 77 | private 78 | def texElabMetavar : TexElab := fun 79 | | _, _, .node _ catName #[.ident _ _ val _] 80 | | _, _, .node _ catName #[.ident _ _ val _, .atom _ "$", .node _ `num _] => do 81 | let ns ← getCurrNamespace 82 | let some qualified := catName.erasePrefix? symbolPrefix | throwUnsupportedSyntax 83 | 84 | let valString := val.toString false 85 | let some { canon, alias, tex? := some tex } := 86 | aliasExt.getState (← getEnv) |>.byAlias.matchPrefix (ns ++ val).toString 0 87 | | return valString.texEscape 88 | if qualified != canon then 89 | return valString.texEscape 90 | let some aliasUnqualified := alias.erasePrefix? ns | return valString.texEscape 91 | let some suffix := valString.dropPrefix? <| aliasUnqualified.toString false 92 | | return valString.texEscape 93 | 94 | return tex ++ suffix.toString.texEscape 95 | | profile, ref, .node _ _ #[«fun», .atom _ "@", idx] => do 96 | let funTex ← texElabMetavar profile ref «fun» 97 | let idxTex ← texElabIdx idx 98 | return s!"{funTex}_\{{idxTex}}" 99 | | _, _, _ => throwUnsupportedSyntax 100 | 101 | mutual 102 | 103 | private partial 104 | def texElabVariable : TexElab := fun 105 | | _, _, .node _ catName #[.node _ _ #[.ident _ _ val _] ] => do 106 | let ns ← getCurrNamespace 107 | let some qualified := catName.erasePrefix? symbolPrefix | throwUnsupportedSyntax 108 | 109 | let valString := val.toString false 110 | let some { canon, alias, tex? := some tex } := 111 | aliasExt.getState (← getEnv) |>.byAlias.matchPrefix (ns ++ val).toString 0 112 | | return valString.texEscape 113 | if qualified != canon then 114 | return valString.texEscape 115 | let some aliasUnqualified := alias.erasePrefix? ns | return valString.texEscape 116 | let some suffix := valString.dropPrefix? <| aliasUnqualified.toString false 117 | | return valString.texEscape 118 | 119 | return tex ++ suffix.toString.texEscape 120 | | profile, ref, .node info kind #[.node _ _ #[«fun», .atom _ "@", idx] ] => do 121 | let funTex ← texElabVariable profile ref <| .node info kind #[«fun»] 122 | let idxTex ← texElabIdx idx 123 | return s!"{funTex}_\{{idxTex}}" 124 | | profile, ref, .node _ catName #[ 125 | base@(.node _ catName₀ _), 126 | .atom _ "[", 127 | val@(.node _ symbolPrefixedValName _), 128 | .atom _ "/", 129 | var@(.node _ symbolPrefixedVarName _), 130 | .atom _ "]" 131 | ] => do 132 | if catName != catName₀ then 133 | throwUnsupportedSyntax 134 | 135 | if !symbolPrefix.isPrefixOf catName then throwUnsupportedSyntax 136 | if !symbolPrefix.isPrefixOf symbolPrefixedValName then throwUnsupportedSyntax 137 | if !symbolPrefix.isPrefixOf symbolPrefixedVarName then throwUnsupportedSyntax 138 | 139 | let baseTex ← texElabSymbolOrJudgement catName profile ref base 140 | let valTex ← texElabSymbolOrJudgement symbolPrefixedValName profile ref val 141 | let varTex ← texElabSymbolOrJudgement symbolPrefixedVarName profile ref var 142 | 143 | return s!"{baseTex}\\left[{valTex}/{varTex}\\right]" 144 | | profile, ref, .node _ catName #[ 145 | base@(.node _ catName₀ _), 146 | .atom _ "^", 147 | var@(.node _ symbolPrefixedVarName _), 148 | .node _ `null level, 149 | node _ `null locallyNamelessSubstAlternative 150 | ] => do 151 | if catName != catName₀ then 152 | throwUnsupportedSyntax 153 | 154 | if !symbolPrefix.isPrefixOf catName then throwUnsupportedSyntax 155 | if !symbolPrefix.isPrefixOf symbolPrefixedVarName then throwUnsupportedSyntax 156 | 157 | let baseTex ← texElabSymbolOrJudgement catName profile ref base 158 | let varTex ← texElabSymbolOrJudgement symbolPrefixedVarName profile ref var 159 | 160 | if !(← getOptions).get lott.tex.locallyNameless.name lott.tex.locallyNameless.defValue then 161 | match locallyNamelessSubstAlternative with 162 | | #[.atom _ "/", var@(.node _ symbolPrefixedVarName _)] => 163 | if !symbolPrefix.isPrefixOf symbolPrefixedVarName then 164 | throwUnsupportedSyntax 165 | 166 | let oldVarTex ← texElabSymbolOrJudgement symbolPrefixedVarName profile ref var 167 | return s!"{baseTex}\\left[{varTex}/{oldVarTex}\\right]" 168 | | #[] => return baseTex 169 | | _ => throwUnsupportedSyntax 170 | 171 | if let some level := level[1]? then 172 | let levelTex ← texElabIdx level 173 | return s!"\{{baseTex}}^\{{varTex}#{levelTex}}" 174 | else 175 | return s!"\{{baseTex}}^\{{varTex}}" 176 | | profile, ref, .node _ catName #[ 177 | base@(.node _ catName₀ _), 178 | .atom _ "^^", 179 | val@(.node _ symbolPrefixedValName _), 180 | .node _ `null level, 181 | altRef@(.node _ `null locallyNamelessSubstAlternative) 182 | ] => do 183 | if catName != catName₀ then 184 | throwUnsupportedSyntax 185 | 186 | if !symbolPrefix.isPrefixOf catName then throwUnsupportedSyntax 187 | if !symbolPrefix.isPrefixOf symbolPrefixedValName then throwUnsupportedSyntax 188 | 189 | let baseTex ← texElabSymbolOrJudgement catName profile ref base 190 | let valTex ← texElabSymbolOrJudgement symbolPrefixedValName profile ref val 191 | 192 | if !(← getOptions).get lott.tex.locallyNameless.name lott.tex.locallyNameless.defValue then 193 | let varTex ← match locallyNamelessSubstAlternative with 194 | | #[.atom _ "/", var@(.node _ symbolPrefixedVarName _)] => 195 | if !symbolPrefix.isPrefixOf symbolPrefixedVarName then 196 | throwUnsupportedSyntax 197 | 198 | texElabSymbolOrJudgement symbolPrefixedVarName profile ref var 199 | | #[] => 200 | logErrorAt altRef 201 | "variable name must be provided with '/x' for alternative tex substitution rendering when lott.tex.locallyNameless is set to false" 202 | throwUnsupportedSyntax 203 | | _ => throwUnsupportedSyntax 204 | return s!"{baseTex}\\left[{valTex}/{varTex}\\right]" 205 | 206 | if let some level := level[1]? then 207 | let levelTex ← texElabIdx level 208 | return s!"\{{baseTex}}^\{{valTex}#{levelTex}}" 209 | else 210 | return s!"\{{baseTex}}^\{{valTex}}" 211 | | profile, ref, .node _ catName #[ 212 | .atom _ "\\", 213 | var@(.node _ symbolPrefixedVarName _), 214 | .node _ `null level, 215 | .atom _ "^", 216 | base@(.node _ catName₀ _) 217 | ] => do 218 | if catName != catName₀ then 219 | throwUnsupportedSyntax 220 | 221 | if !symbolPrefix.isPrefixOf catName then throwUnsupportedSyntax 222 | if !symbolPrefix.isPrefixOf symbolPrefixedVarName then throwUnsupportedSyntax 223 | 224 | let baseTex ← texElabSymbolOrJudgement catName profile ref base 225 | let varTex ← texElabSymbolOrJudgement symbolPrefixedVarName profile ref var 226 | 227 | if !(← getOptions).get lott.tex.locallyNameless.name lott.tex.locallyNameless.defValue then 228 | return s!"\{{baseTex}}" 229 | 230 | if let some level := level[1]? then 231 | let levelTex ← texElabIdx level 232 | return s!"\\leftidx\{{varTex}#{levelTex}}\{{baseTex}}\{}" 233 | else 234 | return s!"\\leftidx\{{varTex}}\{{baseTex}}\{}" 235 | | profile, ref, .node _ parentCatName #[stx@(.node _ childCatName _)] => do 236 | let some parentQualified := parentCatName.erasePrefix? symbolPrefix | throwUnsupportedSyntax 237 | let some childQualified := childCatName.erasePrefix? symbolPrefix | throwUnsupportedSyntax 238 | let some { parent, .. } := childExt.getState (← getEnv) |>.find? childQualified | throwUnsupportedSyntax 239 | if parent != parentQualified then throwUnsupportedSyntax 240 | 241 | texElabSymbolOrJudgement childCatName profile ref stx 242 | | _, _, _ => throwUnsupportedSyntax 243 | 244 | partial 245 | def texElabSymbolOrJudgement (catName : Name) (profile : Name) (ref stx : Syntax) 246 | : TermElabM String := do 247 | let env ← getEnv 248 | let texPrePost? := do 249 | let qualified ← catName.erasePrefix? symbolPrefix 250 | let { texPrePost?, .. } ← symbolExt.getState env |>.find? qualified 251 | texPrePost? 252 | 253 | let rawTex ← match lottTexElabAttribute.getEntries (← getEnv) catName with 254 | | [] => 255 | try 256 | (try texElabMetavar profile ref stx catch _ => texElabVariable profile ref stx) 257 | catch _ => 258 | throwErrorAt ref 259 | "tex elaboration function for '{catName}' has not been implemented{indentD stx}" 260 | | elabFns => 261 | try texElabWithFns elabFns catch ex => 262 | (try texElabMetavar profile ref stx catch _ => 263 | (try texElabVariable profile ref stx catch _ => throw ex)) 264 | 265 | if let some (texPre, texPost) := texPrePost? then 266 | return s!"{texPre} {rawTex} {texPost}" 267 | else 268 | return rawTex 269 | where 270 | texElabWithFns 271 | | [] => unreachable! 272 | | [elabFn] => elabFn.value profile ref stx 273 | | elabFn :: elabFns => do 274 | try 275 | elabFn.value profile ref stx 276 | catch ex => match ex with 277 | | .internal id _ => 278 | if id == unsupportedSyntaxExceptionId then texElabWithFns elabFns else throw ex 279 | | _ => throw ex 280 | 281 | end 282 | 283 | def elabMutualCommands (cs : Array Command) : CommandElabM Unit := 284 | match cs with 285 | | #[] => pure () 286 | | #[c] => elabCommand c 287 | | cs => do elabCommand <| ← `(mutual $cs* end) 288 | 289 | def resolveSymbol (symbolName : Ident) (allowSuffix := true) : CommandElabM Name := do 290 | let name := symbolName.getId 291 | let state := aliasExt.getState (← getEnv) 292 | 293 | let find (n : Name) := if allowSuffix then 294 | state.byAlias.matchPrefix n.toString 0 295 | else 296 | state.byAlias.find? n.toString 297 | 298 | -- First check if the name itself is present in the state. 299 | match find name with 300 | -- If not, check if it's present when prepending the current namespace. 301 | | none => match find <| (← getCurrNamespace) ++ name with 302 | -- Otherwise, check the namespaces we've opened. 303 | | none => 304 | let names ← resolveOpensNames name state #[] <| ← getOpenDecls 305 | let aliases := names.filterMap find 306 | match aliases.foldl (fun acc a => acc.insert a.canon a) (mkNameMap _) |>.toList with 307 | | [] => throwErrorAt symbolName "unknown lott symbol {symbolName}" 308 | | [(canon, _)] => pure canon 309 | | results => 310 | let results := results.toArray.map Prod.snd 311 | let sortedResults := results.qsort (·.alias.toString.length < ·.alias.toString.length) 312 | match sortedResults.filter (·.alias.toString.length == sortedResults[0]!.alias.toString.length) with 313 | | #[] => unreachable! 314 | | #[{ canon, .. }] => pure canon 315 | | results => throwErrorAt symbolName 316 | "ambiguous lott symbol {symbolName}; found multiple matches: {results.map (·.alias)}" 317 | | some { canon, .. } => pure canon 318 | | some { canon, .. } => pure canon 319 | where 320 | resolveOpensNames name state acc 321 | | [] => pure acc 322 | | .simple ns except :: decls => do 323 | let mut acc := acc 324 | if !except.contains name then 325 | acc := acc.push <| ns ++ name 326 | resolveOpensNames name state acc decls 327 | | .explicit id declName :: decls => do 328 | let mut acc := acc 329 | if id = name then 330 | acc := acc.push declName 331 | resolveOpensNames name state acc decls 332 | 333 | partial 334 | def _root_.Lott.IR.ofStx (stx : TSyntax `stx) (l? : Option Ident := none) : CommandElabM IR := do 335 | let l := l?.getD <| ← mkFreshIdent stx 336 | match stx with 337 | | `(stx| $name:ident) => do 338 | let name' := mkIdentFrom name <| name.getId.getFinal 339 | return mk name' <| .category (← resolveSymbol name) 340 | | `(stx| $s:str) => return mk l <| .atom s.getString 341 | | `(stx| sepBy($[$ps]*, $sep)) => 342 | return mk l <| .sepBy (← ps.mapM ofStx) sep.getString 343 | | `(stx| optional($[$ps]*)) => 344 | return mk l <| .optional <| ← ps.mapM ofStx 345 | | _ => throwUnsupportedSyntax 346 | 347 | partial 348 | def _root_.Lott.IR.ofProdArg (arg : TSyntax ``Lott.prodArg) : CommandElabM IR := do 349 | let `(prodArg| $[$l?:ident:]?$stx:stx) := arg | throwUnsupportedSyntax 350 | ofStx stx l? 351 | 352 | partial 353 | def _root_.Lott.IR.toExprArgs (ir : Array IR) (ids binders : Array Name) 354 | : CommandElabM (TSepArray `term ",") := 355 | ir.filterMapM (β := Term) fun | ir@(mk l _) => do 356 | if (← IR.toType ids binders ir).isNone then 357 | return none 358 | 359 | return some l 360 | 361 | def Syntax.mkSymbolEmbed (stx : Syntax) : Term := 362 | .mk <| mkNode ``Lott.symbolEmbed #[mkAtom "[[", stx, mkAtom "]]"] 363 | 364 | def _root_.Lean.Syntax.mkTypeAscription (stx type : Syntax) : Term := 365 | .mk <| mkNode ``Lean.Parser.Term.typeAscription 366 | #[mkAtom "(", stx, mkAtom ":", mkNullNode #[type], mkAtom ")"] 367 | 368 | def _root_.Lean.Syntax.mkMap (collection pat body : Term) (mapName : Name) : MacroM Term := 369 | `($(collection).$(mkIdent mapName):ident fun $pat => $body) 370 | 371 | partial 372 | def _root_.Lott.IR.toMacroSeqItems (ir : Array IR) (canon : Name) (ids binders : Array Name) 373 | : CommandElabM (TSyntaxArray ``Lean.Parser.Term.doSeqItem) := 374 | ir.filterMapM fun 375 | | mk l (.category _) => do 376 | if binders.contains l.getId then 377 | return none 378 | 379 | `(doSeqItem| let $l := Lott.Syntax.mkSymbolEmbed $l) 380 | | mk _ (.atom _) => return none 381 | | mk l (.sepBy ir sep) => do 382 | let catName := sepByPrefix ++ (← getCurrNamespace) ++ canon ++ l.getId |>.obfuscateMacroScopes 383 | let patternArgs ← IR.toPatternArgs ir 384 | let seqItems ← IR.toMacroSeqItems ir canon ids binders 385 | let exprArgs ← IR.toExprArgs ir ids binders 386 | let go ← mkFreshIdent l 387 | 388 | `(doSeqItem| let $l ← match (Syntax.getArgs $l)[0]? with 389 | | some stx => 390 | let rec $go:ident : Lean.Syntax → Lean.MacroM (Bool × Lean.Syntax.Term) 391 | | Lean.Syntax.node _ $(quote catName) #[ 392 | Lean.Syntax.atom _ "" 401 | ] => do 402 | let (singleton, symbol) ← $go:ident symbol 403 | let mapName := if singleton then `map else `flatMap 404 | return (false, ← Lean.Syntax.mkMap (.mk collection) (.mk pat) (.mk symbol) mapName) 405 | | Lean.Syntax.node _ $(quote catName) #[ 406 | lhs@(Lean.Syntax.node _ $(quote catName) _), 407 | Lean.Syntax.atom _ $(quote sep.trim), 408 | rhs@(Lean.Syntax.node _ $(quote catName) _) 409 | ] => do 410 | let (lhsSingleton, lhs) ← $go:ident lhs 411 | let (rhsSingleton, rhs) ← $go:ident rhs 412 | let combineName := if lhsSingleton then ``List.cons else ``HAppend.hAppend 413 | let rhs := if rhsSingleton then 414 | Lean.Syntax.mkCApp ``List.cons #[rhs, Lean.mkIdent ``List.nil] 415 | else 416 | rhs 417 | return (false, Lean.Syntax.mkCApp combineName #[lhs, rhs]) 418 | | Lean.Syntax.node _ $(quote catName) #[$patternArgs,*] => do 419 | $seqItems* 420 | return (true, ← Lott.IR.foldrProd #[$exprArgs,*]) 421 | | _ => Lean.Macro.throwUnsupported 422 | let (singleton, x) ← $go:ident stx 423 | if singleton then 424 | pure <| Lean.Syntax.mkCApp ``List.cons #[x, Lean.mkIdent ``List.nil] 425 | else 426 | pure x 427 | | none => pure <| mkIdent ``List.nil) 428 | | mk l (.optional ir) => do 429 | let catName := optionalPrefix ++ (← getCurrNamespace) ++ canon ++ l.getId |>.obfuscateMacroScopes 430 | let patternArgs ← IR.toPatternArgs ir 431 | let seqItems ← IR.toMacroSeqItems ir canon ids binders 432 | let exprArgs ← IR.toExprArgs ir ids binders 433 | let go ← mkFreshIdent l 434 | 435 | `(doSeqItem| let $l ← match (Syntax.getArgs $l)[0]? with 436 | | some stx => 437 | let rec $go:ident : Lean.Syntax → Lean.MacroM (Bool × Lean.Syntax.Term) 438 | | Lean.Syntax.node _ $(quote catName) #[ 439 | Lean.Syntax.atom _ "" 446 | ] => do 447 | let (option, symbol) ← $go:ident symbol 448 | let modifierName := if option then ``Option.keepIf else ``Option.someIf 449 | return ( 450 | true, 451 | Lean.Syntax.mkCApp modifierName #[symbol, .mk cond] 452 | ) 453 | | Lean.Syntax.node _ $(quote catName) #[$patternArgs,*] => do 454 | $seqItems* 455 | return (false, ← Lott.IR.foldrProd #[$exprArgs,*]) 456 | | _ => Lean.Macro.throwUnsupported 457 | let (option, x) ← $go:ident stx 458 | if option then 459 | pure x 460 | else 461 | pure <| Lean.Syntax.mkCApp ``Option.some #[x] 462 | | none => pure <| mkIdent ``Option.none) 463 | 464 | partial 465 | def _root_.Lott.IR.toTexSeqItems (ir : Array IR) (canon : Name) 466 | : CommandElabM (TSyntaxArray ``Lean.Parser.Term.doSeqItem) := 467 | ir.mapM fun 468 | | mk l (.category n) => do 469 | `(doSeqItem| 470 | let $l ← Lott.texElabSymbolOrJudgement $(quote <| symbolPrefix ++ n) profile ref $l) 471 | | mk l (.atom s) => 472 | `(doSeqItem| let $l := $(quote <| atomToTex s)) 473 | | mk l (.sepBy ir sep) => do 474 | let catName := sepByPrefix ++ (← getCurrNamespace) ++ canon ++ l.getId |>.obfuscateMacroScopes 475 | let patternArgs ← IR.toPatternArgs ir 476 | let seqItems ← IR.toTexSeqItems ir canon 477 | let joinArgs := IR.toJoinArgs ir 478 | let go ← mkFreshIdent l 479 | 480 | `(doSeqItem| let $l ← match (Syntax.getArgs $l)[0]? with 481 | | some stx => 482 | let rec $go:ident : Lean.Syntax → Lean.Elab.Term.TermElabM String 483 | | Lean.Syntax.node _ $(quote catName) #[ 484 | Lean.Syntax.atom _ "" 493 | ] => do 494 | let symbolTex ← $go:ident symbol 495 | if «notex».getNumArgs == 1 then 496 | return s!"\\lottsepbycomprehension\{{symbolTex}}" 497 | 498 | if let some tex := tex.getArgs[3]? then 499 | let customTex := tex.isStrLit?.getD "" 500 | return s!"\\lottsepbycomprehensioncustom\{{symbolTex}}\{{customTex}}" 501 | 502 | let some patSubstring := 503 | pat.getSubstring? (withLeading := false) (withTrailing := false) 504 | | throwUnsupportedSyntax 505 | let patTex := patSubstring.toString.texEscape 506 | let some collectionSubstring := 507 | collection.getSubstring? (withLeading := false) (withTrailing := false) 508 | | throwUnsupportedSyntax 509 | let collectionTex := collectionSubstring.toString.texEscape 510 | return s!"\\lottsepbycomprehensionpatcollection\{{symbolTex}}\{{patTex}}\{{collectionTex}}" 511 | | Lean.Syntax.node _ $(quote catName) #[ 512 | lhs@(Lean.Syntax.node _ $(quote catName) _), 513 | Lean.Syntax.atom _ $(quote sep.trim), 514 | rhs@(Lean.Syntax.node _ $(quote catName) _) 515 | ] => do 516 | let lhsTex ← $go:ident lhs 517 | let rhsTex ← $go:ident rhs 518 | return s!"{lhsTex} {$(quote <| atomToTex sep)} {rhsTex}" 519 | | Lean.Syntax.node _ $(quote catName) #[$patternArgs,*] => do 520 | $seqItems* 521 | return " ".intercalate [$joinArgs,*] 522 | | _ => Lean.Elab.throwUnsupportedSyntax 523 | $go:ident stx 524 | | none => pure "") 525 | | mk l (.optional ir) => do 526 | let catName := optionalPrefix ++ (← getCurrNamespace) ++ canon ++ l.getId |>.obfuscateMacroScopes 527 | let patternArgs ← IR.toPatternArgs ir 528 | let seqItems ← IR.toTexSeqItems ir canon 529 | let joinArgs := IR.toJoinArgs ir 530 | let go ← mkFreshIdent l 531 | 532 | `(doSeqItem| let $l ← match (Syntax.getArgs $l)[0]? with 533 | | some stx => 534 | let rec $go:ident : Lean.Syntax → Lean.Elab.Term.TermElabM String 535 | | Lean.Syntax.node _ $(quote catName) #[ 536 | Lean.Syntax.atom _ "" 543 | ] => do 544 | let symbolTex ← $go:ident symbol 545 | if «notex».getNumArgs == 1 then 546 | return s!"\\lottoptionalcomprehension\{{symbolTex}}" 547 | 548 | if let some tex := tex.getArgs[3]? then 549 | let customTex := tex.isStrLit?.getD "" 550 | return s!"\\lottoptionalcomprehensioncustom\{{symbolTex}}\{{customTex}}" 551 | 552 | let some condSubstring := 553 | cond.getSubstring? (withLeading := false) (withTrailing := false) 554 | | throwUnsupportedSyntax 555 | let condTex := condSubstring.toString.texEscape 556 | return s!"\\lottoptionalcomprehensioncond\{{symbolTex}}\{{condTex}}" 557 | | Lean.Syntax.node _ $(quote catName) #[$patternArgs,*] => do 558 | $seqItems* 559 | return " ".intercalate [$joinArgs,*] 560 | | _ => Lean.Elab.throwUnsupportedSyntax 561 | $go:ident stx 562 | | none => pure "") 563 | where 564 | atomToTex (atom : String) := 565 | let leadingWs := atom.data[0]?.map Char.isWhitespace |>.getD false 566 | let trailingWs := atom.data.getLast?.map Char.isWhitespace |>.getD false 567 | let tex := if atom.trim.data.all (fun c => c.isAlpha || c.isSubscriptAlpha) then 568 | s!"\\lottkw\{{atom.trim.texEscape}}" 569 | else 570 | s!"\\lottsym\{{atom.trim.texEscape}}" 571 | (if leadingWs then "\\, " else "") ++ tex ++ (if trailingWs then " \\," else "") 572 | 573 | partial 574 | def toExampleSyntax (ir : Array IR) (canonQualified profile : Name) (enclosingSepBys := 0) 575 | : CommandElabM (Array Syntax) := do 576 | let inline := (← getOptions).get lott.tex.example.singleProductionInline.name 577 | lott.tex.example.singleProductionInline.defValue 578 | let «notex» := (← getOptions).get lott.tex.example.comprehensionNoTex.name 579 | lott.tex.example.comprehensionNoTex.defValue 580 | ir.filterMapM fun (mk l ir) => match ir with 581 | | .category n => do 582 | let env ← getEnv 583 | if metaVarExt.getState env |>.contains n then 584 | return some <| sepByIdxStrings.foldl (init := mkNode (symbolPrefix ++ n) #[l]) 585 | (stop := enclosingSepBys) 586 | (mkNode (symbolPrefix ++ n) #[·, mkAtom "@", mkIdent <| .str .anonymous ·]) 587 | 588 | if inline then 589 | if let some { normalProds, .. } := symbolExt.getState env |>.find? n then 590 | if normalProds.size == 1 then 591 | return mkNode (symbolPrefix ++ n) <| 592 | ← toExampleSyntax normalProds.toArray[0]!.snd n profile enclosingSepBys 593 | 594 | return mkNode (symbolPrefix ++ n) #[ 595 | sepByIdxStrings.foldl (init := mkNode (variablePrefix ++ n) #[l]) (stop := enclosingSepBys) 596 | (mkNode (variablePrefix ++ n) #[·, mkAtom "@", mkIdent <| .str .anonymous ·]) 597 | ] 598 | | .atom s => return if s == "" then none else mkAtom s.trim 599 | | .sepBy ir _ => do 600 | let catName := sepByPrefix ++ canonQualified ++ l.getId |>.obfuscateMacroScopes 601 | let some patString := sepByIdxStrings[enclosingSepBys]? 602 | | throwError "encountered too many nested sepBys; ran out of index names" 603 | let collectionString := "[:n]" 604 | return mkNullNode #[ 605 | mkNode catName #[ 606 | mkAtom ".toArray, 625 | mkAtom "/>" 626 | ] 627 | ] 628 | | .optional ir => do 629 | let catName := optionalPrefix ++ canonQualified ++ l.getId |>.obfuscateMacroScopes 630 | let condString := "b" 631 | return mkNullNode #[ 632 | mkNode catName #[ 633 | mkAtom ".toArray, 640 | mkAtom "/>" 641 | ] 642 | ] 643 | where 644 | sepByIdxStrings := #["i", "j", "k", "l", "m"] 645 | 646 | /- Conditional syntax. -/ 647 | 648 | elab_rules : command | `(termonly $c) => do if ← getTerm then elabCommand c 649 | 650 | /- Term embedding syntax. -/ 651 | 652 | @[macro qualifiedSymbolEmbed] 653 | def qualifiedSymbolEmbedImpl : Macro := fun stx => 654 | return Lott.Syntax.mkSymbolEmbed <| stx.getArg 1 |>.getArg 2 655 | 656 | end Lott 657 | -------------------------------------------------------------------------------- /Lott/Elab/NonTerminal.lean: -------------------------------------------------------------------------------- 1 | import Lott.Elab.Basic 2 | import Lott.Parser 3 | 4 | namespace Lott 5 | 6 | open Lean 7 | open Lean.Elab 8 | open Lean.Elab.Command 9 | open Lean.Elab.Term 10 | open Lean.Parser 11 | open Lean.Parser.Command 12 | open Lean.Parser.Term 13 | open Lott.IR 14 | 15 | private 16 | def symbolEmbedIdent := mkIdent ``Lott.symbolEmbed 17 | 18 | private partial 19 | def IR.toIsChildCtor (prodIdent isIdent : Ident) (qualified pqualified : Name) (ir pir : Array IR) 20 | (binders : Array Name) : CommandElabM (TSyntax ``Lean.Parser.Command.ctor) := do 21 | if ir.size != pir.size then 22 | throwErrorAt prodIdent "length of child production ({ir.size}) doesn't match length of parent production ({pir.size})" 23 | 24 | let (ctorTypeRetArgs, ctorTypeArgs) ← go (ir.zip pir) 25 | let binders ← ctorTypeRetArgs.filterMapM fun 26 | | `(Lean.binderIdent| $_:hole) => return none 27 | | `(Lean.binderIdent| $i:ident) => `(bracketedBinder| {$i}) 28 | | _ => throwUnsupportedSyntax 29 | let ctorType ← foldrArrow ctorTypeArgs <| ← ``($isIdent ($(mkIdent <| pqualified ++ prodIdent.getId) $(← toTerms ctorTypeRetArgs)*)) 30 | return ← `(ctor| | $prodIdent:ident $binders:bracketedBinder* : $ctorType) 31 | where 32 | go (irs : Array (IR × IR)) (patAcc : TSyntaxArray ``Lean.binderIdent := #[]) 33 | (propAcc : Array Term := #[]) 34 | : CommandElabM (TSyntaxArray ``Lean.binderIdent × Array Term) := do 35 | let some (mk l' ir, mk l pir) := irs[0]? | return (patAcc, propAcc) 36 | let irs' := irs.extract 1 irs.size 37 | 38 | if binders.contains l'.getId then 39 | return ← go irs' patAcc propAcc 40 | 41 | let lbi ← `(Lean.binderIdent| $l:ident) 42 | let l'bi ← `(Lean.binderIdent| $l':ident) 43 | let hole ← `(Lean.binderIdent| _) 44 | 45 | match ir, pir with 46 | | .category n, .category np => do 47 | if n == np then 48 | return ← go irs' (patAcc.push hole) propAcc 49 | 50 | if some np == ((childExt.getState (← getEnv)).find? n).map (·.parent) then 51 | let .str _ nStr := n | throwUnsupportedSyntax 52 | let isName := np ++ Name.str .anonymous nStr |>.appendBefore "Is" 53 | return ← go irs' (patAcc.push lbi) <| propAcc.push <| ← `($(mkIdent isName) $l) 54 | 55 | throwErrorAt prodIdent "couldn't find parent/child relationship between {pqualified} and {qualified}" 56 | | .atom s, .atom sp => do 57 | if s != sp then 58 | throwErrorAt prodIdent "mismatched atoms \"{s}\" and \"{sp}\"" 59 | 60 | go irs' patAcc propAcc 61 | | .sepBy ir sep, .sepBy pir psep => do 62 | if sep != psep then 63 | throwErrorAt prodIdent "mismatched separators \"{sep}\" and \"{psep}\"" 64 | 65 | if ir.size != pir.size then 66 | throwErrorAt prodIdent "length of child sepBy ({ir.size}) doesn't match length of parent sepBy ({pir.size})" 67 | 68 | match ← go (ir.zip pir) with 69 | -- In this case, the sepBy doesn't actually contain anything stored in the datatype so we can 70 | -- just skip it. 71 | | (#[], #[]) => go irs' patAcc propAcc 72 | -- In this case, there is a list with stuff in it, but it doesn't contain anything for us to 73 | -- check, so we can just add an underscore pattern argument and proceed. 74 | | (_, #[]) => go irs' (patAcc.push hole) propAcc 75 | | (#[patArg], props) => go irs' (patAcc.push lbi) <| propAcc.push <| ← 76 | ``(∀ $patArg ∈ $l, $(← foldlAnd props)) 77 | | (patArgs, props) => 78 | let prod ← liftMacroM <| foldrProd <| ← toTerms patArgs 79 | go irs' (patAcc.push lbi) <| propAcc.push <| ← 80 | ``(∀ $l'bi ∈ $l, let $prod:term := $l'; $(← foldlAnd props)) 81 | | .optional ir, .optional pir => do 82 | if ir.size != pir.size then 83 | throwErrorAt prodIdent "length of child optional ({ir.size}) doesn't match length of parent optional ({pir.size})" 84 | 85 | match ← go (ir.zip pir) with 86 | -- In this case, the optional doesn't actually contain anything stored in the datatype so we 87 | -- can just skip it. 88 | | (#[], #[]) => go irs' patAcc propAcc 89 | -- In this case, there is a list with stuff in it, but it doesn't contain anything for us to 90 | -- check, so we can just add an underscore pattern argument and proceed. 91 | | (_, #[]) => go irs' (patAcc.push hole) propAcc 92 | | (#[patArg], props) => go irs' (patAcc.push lbi) <| propAcc.push <| ← 93 | ``(∀ $patArg:binderIdent ∈ $l, $(← foldlAnd props)) 94 | | (patArgs, props) => 95 | let prod ← liftMacroM <| foldrProd <| ← toTerms patArgs 96 | go irs' (patAcc.push lbi) <| propAcc.push <| ← 97 | ``(∀ $l'bi ∈ $l, let $prod:term := $l'; $(← foldlAnd props)) 98 | | _, _ => throwErrorAt prodIdent "mismatched syntax" 99 | 100 | private 101 | def freeName (varName : Name) (typeName : Name) : CommandElabM Name := do 102 | let ns ← getCurrNamespace 103 | let typeName := typeName.replacePrefix ns .anonymous 104 | match varName.replacePrefix ns .anonymous with 105 | | .str .anonymous s => return typeName ++ Name.appendAfter `free s |>.appendAfter "s" 106 | | n => return typeName ++ `free ++ n.appendAfter "s" 107 | 108 | private partial 109 | def toFreeMatchAlt (prodId : Ident) (ir : Array IR) (subst : Name × Name) 110 | (binders : Array Name) : CommandElabM (TSyntax ``Lean.Parser.Term.matchAltExpr) := do 111 | let (patArgs, appendArgs) ← goMany ir 112 | let rhs ← foldlAppend appendArgs 113 | `(matchAltExpr| | $prodId $patArgs* => $rhs) 114 | where 115 | goMany (ir : Array IR) : CommandElabM (Array Term × Array Term) := do 116 | let (patArg?s, appendArg?s) ← Array.unzip <$> ir.mapM go 117 | return (patArg?s.filterMap id, appendArg?s.filterMap id) 118 | go (ir : IR) : CommandElabM (Option Term × Option Term) := do 119 | let .mk l irt := ir 120 | 121 | if binders.contains l.getId then 122 | return (none, none) 123 | 124 | match irt with 125 | | .category n => do 126 | if n == subst.fst then 127 | return (l, some <| ← ``(match $l:ident with | .free x => [x] | .bound _ => [])) 128 | if let some { substitutions, .. } := symbolExt.getState (← getEnv) |>.find? n then 129 | if substitutions.contains subst then 130 | let freeId := mkIdent <| ← freeName subst.fst n 131 | return (l, some <| ← ``($freeId $l)) 132 | return (← ``(_), none) 133 | | .atom _ => return (none, none) 134 | | .sepBy ir _ => match ← goMany ir with 135 | | (#[], _) => return (none, none) 136 | | (_, #[]) => return (← ``(_), none) 137 | | (patArgs, appendArgs) => 138 | let prod ← liftMacroM <| foldrProd patArgs 139 | return ( 140 | l, 141 | ← `((List.mapMem $l fun $prod _ => $(← foldlAppend appendArgs)).flatten) 142 | ) 143 | | .optional ir => match ← goMany ir with 144 | | (#[], _) => return (none, none) 145 | | (_, #[]) => return (← ``(_), none) 146 | | (patArgs, appendArgs) => 147 | let prod ← liftMacroM <| foldrProd patArgs 148 | return ( 149 | l, 150 | ← `( 151 | (Option.mapMem $l fun $prod _ => $(← foldlAppend appendArgs)).getD [] 152 | ) 153 | ) 154 | 155 | private 156 | def locallyClosedName (varName : Name) (typeName : Name) : CommandElabM Name := do 157 | let ns ← getCurrNamespace 158 | let typeName := typeName.replacePrefix ns .anonymous 159 | return typeName ++ varName.replacePrefix ns .anonymous |>.appendAfter "LocallyClosed" 160 | 161 | private partial 162 | def containsVar (varName : Name) (ir : Array IR) (ids binders : Array Name) : Bool := 163 | ir.any fun 164 | | mk l (.category n) => n == varName && !(binders.contains l.getId || ids.contains l.getId) 165 | | mk _ (.atom _) => false 166 | | mk _ (.sepBy ir _) 167 | | mk _ (.optional ir) => containsVar varName ir ids binders 168 | 169 | private partial 170 | def toLocallyClosedCtors (prodId locallyClosedId idxId : Ident) (qualified : Name) (ir : Array IR) 171 | (subst : Name × Name) (ids binders : Array Name) 172 | : CommandElabM (Array (TSyntax ``Lean.Parser.Command.ctor)) := do 173 | if containsVar subst.fst ir ids binders then 174 | let #[mk _ (.category _)] := ir | 175 | throwErrorAt prodId "generation of locally closed definitions is not yet supported for complicated productions" 176 | let boundProdId := mkIdentFrom prodId <| prodId.getId.appendAfter "_bound" 177 | let freeProdId := mkIdentFrom prodId <| prodId.getId.appendAfter "_free" 178 | return #[ 179 | ← `(ctor| | $boundProdId:ident {$idxId} {x} : x < $idxId → $locallyClosedId ($fullProdId (.bound x)) $idxId), 180 | ← `(ctor| | $freeProdId:ident {$idxId} : $locallyClosedId ($fullProdId (.free _)) $idxId) 181 | ] 182 | else 183 | let (ctorTypeRetArgs, ctorTypeArgs) ← go ir 184 | let binders ← ctorTypeRetArgs.filterMapM fun 185 | | `(Lean.binderIdent| $_:hole) => return none 186 | | `(Lean.binderIdent| $i:ident) => `(bracketedBinder| {$i}) 187 | | _ => throwUnsupportedSyntax 188 | let ctorType ← foldrArrow ctorTypeArgs <| ← 189 | ``($locallyClosedId ($fullProdId $(← toTerms ctorTypeRetArgs)*) $idxId) 190 | return #[← `(ctor| | $prodId:ident {$idxId} $binders:bracketedBinder* : $ctorType)] 191 | where 192 | fullProdId := mkIdent <| qualified ++ prodId.getId 193 | inc := ir.any fun 194 | | mk l (.category n) => n == subst.fst && binders.contains l.getId 195 | | _ => false 196 | go (irs : Array IR) (patAcc : TSyntaxArray ``Lean.binderIdent := #[]) (propAcc : Array Term := #[]) 197 | : CommandElabM (TSyntaxArray ``Lean.binderIdent × Array Term) := do 198 | let some (mk l ir) := irs[0]? | return (patAcc, propAcc) 199 | let irs := irs.extract 1 irs.size 200 | 201 | if binders.contains l.getId then 202 | return ← go irs patAcc propAcc 203 | 204 | let lbi ← `(Lean.binderIdent| $l:ident) 205 | let hole ← `(Lean.binderIdent| _) 206 | 207 | match ir with 208 | | .category n => do 209 | if let some { substitutions, .. } := symbolExt.getState (← getEnv) |>.find? n then 210 | if substitutions.contains subst then 211 | let locallyClosedId := mkIdent <| ← locallyClosedName subst.fst n 212 | let prop ← if inc then 213 | ``($locallyClosedId $l ($idxId + 1)) 214 | else 215 | ``($locallyClosedId $l $idxId) 216 | go irs (patAcc.push lbi) <| propAcc.push prop 217 | else 218 | go irs (patAcc.push hole) propAcc 219 | else 220 | go irs (patAcc.push hole) propAcc 221 | | .atom _ => go irs patAcc propAcc 222 | | .sepBy ir _ 223 | | .optional ir => do 224 | match ← go ir with 225 | -- In this case, the sepBy or optional doesn't actually contain anything stored in the 226 | -- datatype so we can just skip it. 227 | | (#[], #[]) => go irs patAcc propAcc 228 | -- In this case, there is stuff in it, but it doesn't contain anything for us to check, so we 229 | -- can just add an underscore pattern argument and proceed. 230 | | (_, #[]) => go irs (patAcc.push hole) propAcc 231 | | (#[patArg], props) => go irs (patAcc.push lbi) <| propAcc.push <| ← 232 | `(∀ $patArg:binderIdent ∈ $l, $(← foldlAnd props)) 233 | | (patArgs, props) => 234 | let propAcc' ← props.mapM fun prop => 235 | do `(∀ x ∈ $l, let $(← liftMacroM <| foldrProd <| ← toTerms patArgs):term := x; $prop) 236 | go irs (patAcc.push lbi) <| propAcc ++ propAcc' 237 | 238 | private 239 | def substName (varName : Name) : CommandElabM Name := 240 | return varName.replacePrefix (← getCurrNamespace) .anonymous |>.appendAfter "_subst" 241 | 242 | private partial 243 | def toSubstPatternAndArgs (ir : Array IR) (varType valType : Name) (varId valId : Ident) 244 | (bindOf? binderId? : Option Ident) : CommandElabM (Option (Array Term × Array Term)) := do 245 | let patsAndArgsAndFounds ← ir.filterMapM fun (mk l irt) => do 246 | if some l == bindOf? then return none 247 | 248 | match irt with 249 | | .category n => 250 | if let some { substitutions, .. } := symbolExt.getState (← getEnv) |>.find? n then 251 | if substitutions.contains (varType, valType) then 252 | let substId := mkIdent <| n ++ (← substName varType) 253 | return some (l, ← `($substId:ident $l $varId $valId), true) 254 | 255 | return some (l, l, false) 256 | | .atom _ => return none 257 | | .sepBy ir _ => 258 | let some (pattern, args) ← 259 | toSubstPatternAndArgs ir varType valType varId valId bindOf? binderId? | 260 | return some (l, (l : Term), false) 261 | let patProd ← liftMacroM <| foldrProd pattern 262 | let argsProd ← liftMacroM <| foldrProd args 263 | return some (l, ← `(List.mapMem $l fun $patProd _ => $argsProd), true) 264 | | .optional ir => 265 | let some (pattern, args) ← 266 | toSubstPatternAndArgs ir varType valType varId valId bindOf? binderId? | 267 | return some (l, (l : Term), false) 268 | let patProd ← liftMacroM <| foldrProd pattern 269 | let argsProd ← liftMacroM <| foldrProd args 270 | let arg ← `(Option.mapMem $l fun $patProd _ => $argsProd) 271 | return some (l, arg, true) 272 | let (pats, argsAndFounds) := patsAndArgsAndFounds.unzip 273 | let (args, founds) := argsAndFounds.unzip 274 | if founds.any id then 275 | return some (pats, args) 276 | else 277 | return none 278 | 279 | private 280 | def openName (varName : Name) : CommandElabM Name := 281 | return varName.replacePrefix (← getCurrNamespace) .anonymous |>.appendAfter "_open" 282 | 283 | private partial 284 | def toVarOpenPatternAndArgs (ir : Array IR) (varType valType : Name) (varId idxId : Ident) 285 | (bindOf? binderId? : Option Ident) : CommandElabM (Option (Array Term × Array Term)) := do 286 | let patsAndArgsAndFounds ← ir.filterMapM fun (mk l irt) => do 287 | if some l == bindOf? then return none 288 | 289 | match irt with 290 | | .category n => 291 | if some l != binderId? && n == varType then 292 | return some (l, ← `(if .bound $idxId = $l then .free $varId else $l), true) 293 | 294 | if let some { substitutions, .. } := symbolExt.getState (← getEnv) |>.find? n then 295 | if substitutions.contains (varType, valType) then 296 | return some (l, ← `($(mkIdent <| n ++ (← openName varType)):ident $l $varId $idxId), true) 297 | 298 | return some (l, l, false) 299 | | .atom _ => return none 300 | | .sepBy ir _ => 301 | let some (pattern, args) ← 302 | toVarOpenPatternAndArgs ir varType valType varId idxId bindOf? binderId? | 303 | return some (l, (l : Term), false) 304 | let patProd ← liftMacroM <| foldrProd pattern 305 | let argsProd ← liftMacroM <| foldrProd args 306 | return some (l, ← `(List.mapMem $l fun $patProd _ => $argsProd), true) 307 | | .optional ir => 308 | let some (pattern, args) ← 309 | toVarOpenPatternAndArgs ir varType valType varId idxId bindOf? binderId? | 310 | return some (l, (l : Term), false) 311 | let patProd ← liftMacroM <| foldrProd pattern 312 | let argsProd ← liftMacroM <| foldrProd args 313 | let arg ← `(Option.mapMem $l fun $patProd _ => $argsProd) 314 | return some (l, arg, true) 315 | let (pats, argsAndFounds) := patsAndArgsAndFounds.unzip 316 | let (args, founds) := argsAndFounds.unzip 317 | if founds.any id then 318 | return some (pats, args) 319 | else 320 | return none 321 | 322 | private partial 323 | def toValOpenPatternAndArgs (ir : Array IR) (varType valType : Name) (valId idxId : Ident) 324 | (bindOf? binderId? : Option Ident) : CommandElabM (Option (Array Term × Array Term)) := do 325 | let patsAndArgsAndFounds ← ir.filterMapM fun (mk l irt) => do 326 | if some l == bindOf? then return none 327 | 328 | match irt with 329 | | .category n => 330 | if let some { substitutions, .. } := symbolExt.getState (← getEnv) |>.find? n then 331 | if substitutions.contains (varType, valType) then 332 | return some (l, ← `($(mkIdent <| n ++ (← openName valType)):ident $l $valId $idxId), true) 333 | 334 | return some (l, l, false) 335 | | .atom _ => return none 336 | | .sepBy ir _ => 337 | let some (pattern, args) ← 338 | toValOpenPatternAndArgs ir varType valType valId idxId bindOf? binderId? | 339 | return some (l, (l : Term), false) 340 | let patProd ← liftMacroM <| foldrProd pattern 341 | let argsProd ← liftMacroM <| foldrProd args 342 | return some (l, ← `(List.mapMem $l fun $patProd _ => $argsProd), true) 343 | | .optional ir => 344 | let some (pattern, args) ← 345 | toValOpenPatternAndArgs ir varType valType valId idxId bindOf? binderId? | 346 | return some (l, (l : Term), false) 347 | let patProd ← liftMacroM <| foldrProd pattern 348 | let argsProd ← liftMacroM <| foldrProd args 349 | let arg ← `(Option.mapMem $l fun $patProd _ => $argsProd) 350 | return some (l, arg, true) 351 | let (pats, argsAndFounds) := patsAndArgsAndFounds.unzip 352 | let (args, founds) := argsAndFounds.unzip 353 | if founds.any id then 354 | return some (pats, args) 355 | else 356 | return none 357 | 358 | private 359 | def closeName (varName : Name) : CommandElabM Name := 360 | return varName.replacePrefix (← getCurrNamespace) .anonymous |>.appendAfter "_close" 361 | 362 | private partial 363 | def toClosePatternAndArgs (ir : Array IR) (varType valType : Name) (varId idxId : Ident) 364 | (bindOf? binderId? : Option Ident) : CommandElabM (Option (Array Term × Array Term)) := do 365 | let patsAndArgsAndFounds ← ir.filterMapM fun (mk l irt) => do 366 | if some l == bindOf? then return none 367 | 368 | match irt with 369 | | .category n => 370 | if some l != binderId? && n == varType then 371 | return some (l, ← `(if .free $varId = $l then .bound $idxId else $l), true) 372 | 373 | if let some { substitutions, .. } := symbolExt.getState (← getEnv) |>.find? n then 374 | if substitutions.contains (varType, valType) then 375 | return some (l, ← `($(mkIdent <| n ++ (← closeName varType)):ident $l $varId $idxId), true) 376 | 377 | return some (l, l, false) 378 | | .atom _ => return none 379 | | .sepBy ir _ => 380 | let some (pattern, args) ← 381 | toClosePatternAndArgs ir varType valType varId idxId bindOf? binderId? | 382 | return some (l, (l : Term), false) 383 | let patProd ← liftMacroM <| foldrProd pattern 384 | let argsProd ← liftMacroM <| foldrProd args 385 | return some (l, ← `(List.mapMem $l fun $patProd _ => $argsProd), true) 386 | | .optional ir => 387 | let some (pattern, args) ← 388 | toClosePatternAndArgs ir varType valType varId idxId bindOf? binderId? | 389 | return some (l, (l : Term), false) 390 | let patProd ← liftMacroM <| foldrProd pattern 391 | let argsProd ← liftMacroM <| foldrProd args 392 | let arg ← `(Option.mapMem $l fun $patProd _ => $argsProd) 393 | return some (l, arg, true) 394 | let (pats, argsAndFounds) := patsAndArgsAndFounds.unzip 395 | let (args, founds) := argsAndFounds.unzip 396 | if founds.any id then 397 | return some (pats, args) 398 | else 399 | return none 400 | 401 | @[macro symbolEmbed] 402 | private 403 | def nonTerminalImpl : Macro := fun 404 | | .node _ ``Lott.symbolEmbed #[ 405 | .atom _ "[[", 406 | .node _ _ #[.node _ _ #[ident@(.ident ..)] ], 407 | .atom _ "]]" 408 | ] => return ident 409 | | .node _ ``Lott.symbolEmbed #[ 410 | .atom _ "[[", 411 | .node info kind #[.node _ _ #[«fun», .atom _ "@", idx] ], 412 | .atom _ "]]" 413 | ] => `(($(Syntax.mkSymbolEmbed (.node info kind #[«fun»])) : _ → _) $(.mk idx)) 414 | -- NOTE: Could maybe make these even more resistant to false positives on user definitions by 415 | -- adding special intermediate nodes; not really any need though unless people actually run into 416 | -- problems. 417 | | .node _ ``Lott.symbolEmbed #[ 418 | .atom _ "[[", 419 | .node _ catName #[ 420 | base@(.node _ catName₀ _), 421 | .atom _ "[", 422 | val@(.node _ symbolPrefixedValName _), 423 | .atom _ "/", 424 | var@(.node _ symbolPrefixedVarName _), 425 | .atom _ "]" 426 | ], 427 | .atom _ "]]" 428 | ] => do 429 | if catName != catName₀ then 430 | Macro.throwUnsupported 431 | 432 | let some canonQualified := catName.erasePrefix? symbolPrefix | Macro.throwUnsupported 433 | if !symbolPrefix.isPrefixOf symbolPrefixedValName then Macro.throwUnsupported 434 | let some varName := symbolPrefixedVarName.erasePrefix? symbolPrefix | Macro.throwUnsupported 435 | 436 | let (_, varName) := canonQualified.removeCommonPrefixes varName 437 | let substName := canonQualified ++ varName.appendAfter "_subst" 438 | 439 | return .mkCApp substName #[ 440 | Syntax.mkSymbolEmbed base, 441 | Syntax.mkSymbolEmbed var, 442 | Syntax.mkSymbolEmbed val 443 | ] 444 | | .node _ ``Lott.symbolEmbed #[ 445 | .atom _ "[[", 446 | .node _ catName #[ 447 | base@(.node _ catName₀ _), 448 | .atom _ "^", 449 | var@(.node _ symbolPrefixedVarName _), 450 | .node _ `null level, 451 | .node _ `null _ 452 | ], 453 | .atom _ "]]" 454 | ] => do 455 | if catName != catName₀ then 456 | Macro.throwUnsupported 457 | 458 | let some canonQualified := catName.erasePrefix? symbolPrefix | Macro.throwUnsupported 459 | let some varName := symbolPrefixedVarName.erasePrefix? symbolPrefix | Macro.throwUnsupported 460 | 461 | let (_, varName) := canonQualified.removeCommonPrefixes varName 462 | let openName := canonQualified ++ varName.appendAfter "_open" 463 | 464 | return .mkCApp openName <| #[ 465 | Syntax.mkSymbolEmbed base, 466 | Syntax.mkSymbolEmbed var 467 | ] ++ (level[1]?.map TSyntax.mk).toArray 468 | | .node _ ``Lott.symbolEmbed #[ 469 | .atom _ "[[", 470 | .node _ catName #[ 471 | base@(.node _ catName₀ _), 472 | .atom _ "^^", 473 | val@(.node _ symbolPrefixedValName _), 474 | .node _ `null level, 475 | .node _ `null _ 476 | ], 477 | .atom _ "]]" 478 | ] => do 479 | if catName != catName₀ then 480 | Macro.throwUnsupported 481 | 482 | let some canonQualified := catName.erasePrefix? symbolPrefix | Macro.throwUnsupported 483 | let some valName := symbolPrefixedValName.erasePrefix? symbolPrefix | Macro.throwUnsupported 484 | 485 | let (_, valName) := canonQualified.removeCommonPrefixes valName 486 | let openName := canonQualified ++ valName.appendAfter "_open" 487 | 488 | return .mkCApp openName <| #[ 489 | Syntax.mkSymbolEmbed base, 490 | Syntax.mkSymbolEmbed val 491 | ] ++ (level[1]?.map TSyntax.mk).toArray 492 | | .node _ ``Lott.symbolEmbed #[ 493 | .atom _ "[[", 494 | .node _ catName #[ 495 | .atom _ "\\", 496 | var@(.node _ symbolPrefixedVarName _), 497 | .node _ `null level, 498 | .atom _ "^", 499 | base@(.node _ catName₀ _) 500 | ], 501 | .atom _ "]]" 502 | ] => do 503 | if catName != catName₀ then 504 | Macro.throwUnsupported 505 | 506 | let some canonQualified := catName.erasePrefix? symbolPrefix | Macro.throwUnsupported 507 | let some varName := symbolPrefixedVarName.erasePrefix? symbolPrefix | Macro.throwUnsupported 508 | 509 | let (_, varName) := canonQualified.removeCommonPrefixes varName 510 | let openName := canonQualified ++ varName.appendAfter "_close" 511 | 512 | return .mkCApp openName <| #[ 513 | Syntax.mkSymbolEmbed base, 514 | Syntax.mkSymbolEmbed var 515 | ] ++ (level[1]?.map TSyntax.mk).toArray 516 | | _ => Macro.throwUnsupported 517 | 518 | private 519 | structure BindConfig where 520 | of : Ident 521 | in' : Array Ident 522 | deriving BEq 523 | 524 | private 525 | def BindConfig.find? (bc : BindConfig) (n : Name) : Option Ident := Id.run do 526 | if bc.of.getId == n then 527 | return bc.of 528 | 529 | bc.in'.find? (·.getId == n) 530 | 531 | instance : ToStream BindConfig (Array Ident) where 532 | toStream | { of, in' } => in'.push of 533 | 534 | private 535 | structure Production where 536 | name : Ident 537 | ir : Array IR 538 | expand? : Option Term 539 | bindConfig? : Option BindConfig 540 | ids : Array Name 541 | «notex» : Bool 542 | profileTex : NameMap Term 543 | 544 | private 545 | def Production.binders (prod : Production) : Array Name := 546 | prod.bindConfig?.toArray.map fun { of, .. } => of.getId 547 | 548 | private 549 | structure NonTerminal where 550 | canon : Ident 551 | aliases : Array (Ident × Bool × Option String) 552 | prods : Array Production 553 | parent? : Option Ident 554 | substitutions? : Option (Array (Name × Name)) 555 | texPrePost? : Option (String × String) 556 | profiles : Array Name 557 | 558 | namespace NonTerminal 559 | 560 | private 561 | def qualified (nt : NonTerminal) : CommandElabM Name := 562 | return (← getCurrNamespace) ++ nt.canon.getId 563 | 564 | private 565 | def catName (nt : NonTerminal) : CommandElabM Name := 566 | return symbolPrefix ++ (← nt.qualified) 567 | 568 | private 569 | def varCatName (nt : NonTerminal) : CommandElabM Name := 570 | return variablePrefix ++ (← nt.qualified) 571 | 572 | private 573 | def parserAttrName (nt : NonTerminal) : CommandElabM Name := 574 | return (← nt.catName).appendAfter "_parser" 575 | 576 | private 577 | def varParserAttrName (nt : NonTerminal) : CommandElabM Name := 578 | return (← nt.varCatName).appendAfter "_parser" 579 | 580 | private partial 581 | def shallowSubstitutionClosure (nt : NonTerminal) (qualifiedLocalMap : NameMap NonTerminal) 582 | : CommandElabM NameSet := do 583 | let mut res := .empty |>.insert (← nt.qualified) 584 | let mut queue := #[nt] 585 | repeat do 586 | let some nt := queue.back? | break 587 | queue := queue.pop 588 | 589 | for { ir, .. } in nt.prods do 590 | for name in irsClosure ir do 591 | if res.contains name then 592 | continue 593 | 594 | res := res.insert name 595 | 596 | let some nt := qualifiedLocalMap.find? name | continue 597 | queue := queue.push nt 598 | return res 599 | where 600 | irClosure : IR → NameSet 601 | | .mk _ (.category n) => .empty |>.insert n 602 | | .mk _ (.atom _) => .empty 603 | | .mk _ (.sepBy ir _) 604 | | .mk _ (.optional ir) => irsClosure ir 605 | irsClosure ir := ir.foldl (·.union <| irClosure ·) .empty 606 | 607 | private partial 608 | def profileClosure (nt : NonTerminal) (qualifiedLocalMap : NameMap NonTerminal) 609 | : CommandElabM (Array Name) := do 610 | let mut res : NameSet := RBTree.fromArray nt.profiles Name.quickCmp 611 | let mut visited : NameSet := .empty |>.insert <| ← nt.qualified 612 | let mut queue := #[nt] 613 | repeat do 614 | let some nt := queue.back? | break 615 | queue := queue.pop 616 | 617 | for { ir, .. } in nt.prods do 618 | for name in irsClosure ir do 619 | if visited.contains name then 620 | continue 621 | 622 | visited := visited.insert name 623 | 624 | if let some nt@{ profiles, .. } := qualifiedLocalMap.find? name then 625 | res := res.union <| .fromArray profiles Name.quickCmp 626 | queue := queue.push nt 627 | else if let some { profiles, .. } := symbolExt.getState (← getEnv) |>.find? name then 628 | res := res.union <| .fromArray profiles Name.quickCmp 629 | return res.toArray 630 | where 631 | irClosure : IR → NameSet 632 | | .mk _ (.category n) => .empty |>.insert n 633 | | .mk _ (.atom _) => .empty 634 | | .mk _ (.sepBy ir _) 635 | | .mk _ (.optional ir) => irsClosure ir 636 | irsClosure ir := ir.foldl (·.union <| irClosure ·) .empty 637 | 638 | end NonTerminal 639 | 640 | set_option maxHeartbeats 2000000 in 641 | private 642 | def elabNonTerminals (nts : Array Syntax) : CommandElabM Unit := do 643 | -- Populate alias map immediately so we can resolve stuff within the current mutual throughout the 644 | -- following steps. 645 | let ns ← getCurrNamespace 646 | for nt in nts do 647 | let `(NonTerminal| $[nosubst]? nonterminal $[(parent := $parent?)]? $[(tex pre := $_, post := $_)]? $[$names $[notex]? $[(tex := $nameTex?s)]?],* := $_*) := nt 648 | | throwUnsupportedSyntax 649 | let names@(canonName :: _) := names.toList | throwUnsupportedSyntax 650 | let canon := ns ++ canonName.getId 651 | for (name, tex?) in names.zip nameTex?s.toList do 652 | let tex? := tex?.map TSyntax.getString 653 | setEnv <| aliasExt.addEntry (← getEnv) { canon, alias := ns ++ name.getId, tex? } 654 | 655 | if let some parent := parent? then 656 | let parent ← resolveSymbol parent (allowSuffix := false) 657 | setEnv <| childExt.addEntry (← getEnv) { canon, parent } 658 | 659 | -- Transform syntax into non-terminal structure. 660 | let nts ← nts.mapM fun nt => do 661 | let `(NonTerminal| $[nosubst%$ns]? nonterminal $[(parent := $parent?)]? $[(tex pre := $pre?, post := $post?)]? $[$names $[notex%$nameNt?s]? $[(tex := $nameTex?s)]?],* := $prods*) := nt 662 | | throwUnsupportedSyntax 663 | let some canon := names[0]? | throwUnsupportedSyntax 664 | let aliases := names.extract 1 names.size |>.zip <| 665 | nameNt?s.extract 1 nameNt?s.size |>.map Option.isSome |>.zip <| 666 | nameTex?s.extract 1 nameTex?s.size |>.map <| Option.map TSyntax.getString 667 | 668 | let prodSubstAndProfiles ← prods.mapM fun prod => do 669 | let `(Production| | $[$ps]* : $name $[nosubst%$ns?]? $[notex%$nt?]? $[(bind $of? $[in $in??,*]?)]? $[(id $ids?,*)]? $[(expand := $expand?)]? $[(tex $[$texProfile?s]? := $tex?s)]*) := prod 670 | | throwUnsupportedSyntax 671 | let ids := ids?.getD (.mk #[]) |>.getElems 672 | 673 | let ir ← ps.mapM IR.ofProdArg 674 | let subst ← match ir with 675 | | #[mk _ (.category n)] => 676 | if (metaVarExt.getState (← getEnv)).contains n then 677 | pure <| if ns?.isNone then some n else none 678 | else 679 | if let some ns := ns? then 680 | logWarningAt ns "unused nosubst; production is not a candidate for substitution" 681 | pure none 682 | | _ => 683 | if let some ns := ns? then 684 | logWarningAt ns "unused nosubst; production is not a candidate for substitution" 685 | pure none 686 | 687 | let bindConfig? ← match of?, in?? with 688 | | some of, some in? => do 689 | let res := { of, in' := in?.getD (.mk #[]) : BindConfig } 690 | if let some x := ids.find? (res.find? ·.getId |>.isSome) then 691 | throwErrorAt x "name {x} also appears in bind config" 692 | for name in toStream res do 693 | if !containsName ir name.getId then 694 | logWarningAt name "name not found in syntax" 695 | pure <| some res 696 | | none, none => pure none 697 | | _, _ => unreachable! 698 | 699 | let ids := ids.map TSyntax.getId 700 | 701 | let (_, defaults) := texProfile?s.zip tex?s |>.partition fun (p?, _) => p?.isSome 702 | if defaults.size > 1 then 703 | throwUnsupportedSyntax 704 | let profileTex := RBMap.fromArray (cmp := Name.quickCmp) <| texProfile?s.zip tex?s |>.map 705 | fun (profile?, tex) => (profile?.map TSyntax.getId |>.getD default, tex) 706 | 707 | return ( 708 | { name, ir, expand?, bindConfig?, ids, «notex» := nt?.isSome, profileTex }, 709 | subst, 710 | texProfile?s.filterMap id 711 | ) 712 | 713 | let (prods, substAndProfiles) := prodSubstAndProfiles.unzip 714 | let (substs, profiless) := substAndProfiles.unzip 715 | let qualified := (← getCurrNamespace) ++ canon.getId 716 | let substitutions? := if parent?.isSome || ns.isSome then 717 | none 718 | else 719 | substs.filterMap id |>.map ((·, qualified)) 720 | 721 | let parent? ← parent?.mapM fun parent => 722 | return mkIdentFrom parent <| ← resolveSymbol parent (allowSuffix := false) 723 | 724 | let texPrePost? := match pre?, post? with 725 | | some pre, some post => some (pre.getString, post.getString) 726 | | none, none => none 727 | | _, _ => unreachable! 728 | 729 | let profiles := RBTree.toArray <| RBTree.fromArray (cmp := Name.quickCmp) <| 730 | profiless.flatten.map TSyntax.getId 731 | 732 | return { canon, aliases, prods, parent?, substitutions?, texPrePost?, profiles : NonTerminal } 733 | 734 | -- Define parser categories. 735 | for nt in nts do 736 | setEnv <| ← registerParserCategory (← getEnv) (← nt.parserAttrName) (← nt.catName) .symbol 737 | setEnv <| ← registerParserCategory (← getEnv) (← nt.varParserAttrName) (← nt.varCatName) .symbol 738 | 739 | -- Define mutual inductives. 740 | if ← getTerm then 741 | let defsAndInductives ← nts.mapM fun nt@{ canon, prods, parent?, .. } => do 742 | let some parent := parent? | 743 | let ctors ← prods.filterMapM fun prod@{ name, ir, expand?, ids, .. } => do 744 | let none := expand? | return none 745 | let ctorType ← IR.toTypeArrSeq ir canon ids prod.binders 746 | `(ctor| | $name:ident : $ctorType) 747 | let inductive' ← `(inductive $canon where $ctors*) 748 | return (none, inductive') 749 | let some { normalProds, .. } := symbolExt.getState (← getEnv) |>.find? parent.getId 750 | | throwErrorAt parent "unknown parent {parent.getId}" 751 | 752 | let isIdent := mkIdentFrom canon <| parent.getId ++ canon.getId.appendBefore "Is" 753 | let ctors ← prods.mapM fun prod@{ name, ir, expand?, .. } => match expand? with 754 | | some ref => 755 | throwErrorAt ref "nonterminal with parent cannot contain productions with expand specifier" 756 | | none => do 757 | let some pir := normalProds.find? name.getId 758 | | throwErrorAt name "normal production with matching name not found in parent" 759 | 760 | IR.toIsChildCtor name isIdent (← nt.qualified) parent.getId ir pir prod.binders 761 | let inductive' ← 762 | `(inductive $(mkIdentFrom isIdent <| `_root_ ++ isIdent.getId) : $parent → Prop where 763 | $ctors*) 764 | 765 | return (some <| ← `(def $canon := { x : $parent // $isIdent x }), inductive') 766 | let (defs, inductives) := defsAndInductives.unzip 767 | elabMutualCommands inductives 768 | elabMutualCommands <| defs.filterMap id 769 | else 770 | for { canon, .. } in nts do 771 | elabCommand <| ← `(opaque $canon : Type) 772 | 773 | let ntsMap ← nts.foldlM (init := mkNameMap _) fun acc nt => return acc.insert (← nt.qualified) nt 774 | let namePairCmp | (a₁, a₂), (b₁, b₂) => Name.quickCmp a₁ b₁ |>.«then» <| Name.quickCmp a₂ b₂ 775 | let nts ← nts.mapM fun nt => do 776 | if nt.substitutions?.isNone then 777 | return nt 778 | 779 | let mut substitutions' := RBTree.empty (cmp := namePairCmp) 780 | for n in ← nt.shallowSubstitutionClosure ntsMap do 781 | let substitutions? ← match ntsMap.find? n with 782 | | some { substitutions?, .. } => pure substitutions? 783 | | none => 784 | let some { substitutions, .. } := symbolExt.getState (← getEnv) |>.find? n | continue 785 | pure substitutions 786 | let some substitutions := substitutions? | continue 787 | for subst in substitutions do 788 | substitutions' := substitutions'.insert subst 789 | return { nt with substitutions? := substitutions'.toArray : NonTerminal } 790 | 791 | let nts ← nts.mapM fun nt => do 792 | return { nt with profiles := ← nt.profileClosure ntsMap : NonTerminal } 793 | 794 | for nt@{ canon, aliases, prods, parent?, substitutions?, texPrePost?, profiles } in nts do 795 | -- Add symbol to environment extension before proceeding so that lookups of things within the 796 | -- current mutual still work properly. 797 | let normalProds := prods.foldl (init := mkNameMap _) fun acc { name, ir, expand?, .. } => 798 | if expand?.isNone then acc.insert name.getId ir else acc 799 | setEnv <| symbolExt.addEntry (← getEnv) { 800 | qualified := ← nt.qualified, 801 | normalProds, 802 | substitutions := substitutions?.getD #[], 803 | texPrePost?, 804 | profiles 805 | } 806 | 807 | let isLocallyNameless (varName : Name) : CommandElabM Bool := 808 | return metaVarExt.getState (← getEnv) |>.find! varName 809 | 810 | for nt@{ canon, aliases, prods, parent?, substitutions?, .. } in nts do 811 | -- Define production and substitution parsers. 812 | let canonName := canon.getId 813 | let attrIdent := mkIdent <| ← nt.parserAttrName 814 | if parent?.isNone then 815 | for { name, ir, .. } in prods do 816 | let (val, type) ← IR.toParser ir nt.canon.getId symbolPrefix 817 | let parserIdent := mkIdentFrom name <| canonName ++ name.getId.appendAfter "_parser" 818 | elabCommand <| ← 819 | `(@[Lott.Symbol_parser, $attrIdent:ident] 820 | private 821 | def $parserIdent : $type := $val) 822 | 823 | let substitutions := substitutions?.getD #[] 824 | for (varName, valName) in substitutions do 825 | let parserIdent := mkIdent <| canonName ++ varName.appendAfter "_subst_parser" 826 | elabCommand <| ← 827 | `(@[Lott.Symbol_parser, $attrIdent:ident] 828 | private 829 | def $parserIdent : TrailingParser := 830 | trailingNode $(quote <| ← nt.catName) Parser.maxPrec 0 <| 831 | "[" >> categoryParser $(quote <| symbolPrefix ++ valName) 0 >> " / " >> 832 | categoryParser $(quote <| symbolPrefix ++ varName) 0 >> "]") 833 | 834 | -- Also add open parsers for this variable if it's locally nameless. 835 | if !(← isLocallyNameless varName) then 836 | continue 837 | 838 | let parserIdent := mkIdent <| canonName ++ varName.appendAfter "_open_parser" 839 | elabCommand <| ← 840 | `(@[Lott.Symbol_parser, $attrIdent:ident] 841 | private 842 | def $parserIdent : TrailingParser := 843 | trailingNode $(quote <| ← nt.catName) Parser.maxPrec 0 <| 844 | checkNoWsBefore >> "^" >> categoryParser $(quote <| symbolPrefix ++ varName) 0 >> 845 | Parser.optional (checkNoWsBefore >> "#" >> checkLineEq >> Parser.numLit) >> 846 | Parser.optional (checkNoWsBefore >> "/" >> checkLineEq >> 847 | categoryParser $(quote <| symbolPrefix ++ varName) 0)) 848 | 849 | let parserIdent := mkIdent <| canonName ++ valName.appendAfter "_open_parser" 850 | elabCommand <| ← 851 | `(@[Lott.Symbol_parser, $attrIdent:ident] 852 | private 853 | def $parserIdent : TrailingParser := 854 | trailingNode $(quote <| ← nt.catName) Parser.maxPrec 0 <| 855 | checkNoWsBefore >> "^^" >> categoryParser $(quote <| symbolPrefix ++ valName) 0 >> 856 | Parser.optional (checkNoWsBefore >> "#" >> checkLineEq >> Parser.numLit) >> 857 | Parser.optional (checkNoWsBefore >> "/" >> checkLineEq >> 858 | categoryParser $(quote <| symbolPrefix ++ varName) 0)) 859 | 860 | let parserIdent := mkIdent <| canonName ++ varName.appendAfter "_close_parser" 861 | elabCommand <| ← 862 | `(@[Lott.Symbol_parser, $attrIdent:ident] 863 | private 864 | def $parserIdent : Parser := 865 | leadingNode $(quote <| ← nt.catName) Parser.maxPrec <| 866 | "\\" >> categoryParser $(quote <| symbolPrefix ++ varName) 0 >> 867 | Parser.optional (checkNoWsBefore >> "#" >> checkLineEq >> Parser.numLit) >> 868 | "^" >> categoryParser $(quote <| ← nt.catName) 0) 869 | 870 | -- Define variable parsers, plus variable category parser in symbol category. 871 | let varAttrIdent := mkIdent <| ← nt.varParserAttrName 872 | for (alias, _) in aliases do 873 | let aliasName@(.str .anonymous nameStr) := alias.getId | throwUnsupportedSyntax 874 | let parserIdent := mkIdentFrom alias <| canonName ++ aliasName.appendAfter "_parser" 875 | elabCommand <| ← 876 | `(@[$varAttrIdent:ident] 877 | private 878 | def $parserIdent : Parser := 879 | leadingNode $(quote <| ← nt.varCatName) Parser.maxPrec <| 880 | Lott.identPrefix $(quote nameStr)) 881 | 882 | let parserIdent := mkIdentFrom canon <| canon.getId.appendAfter "_idx_parser" 883 | elabCommand <| ← 884 | `(@[$varAttrIdent:ident] 885 | private 886 | def $parserIdent : TrailingParser := 887 | trailingNode $(quote <| ← nt.varCatName) Parser.maxPrec 0 <| 888 | checkNoWsBefore >> "@" >> checkLineEq >> (Parser.ident <|> Parser.numLit)) 889 | 890 | let varParserIdent := mkIdentFrom canon <| canonName.appendAfter "_variable_parser" 891 | elabCommand <| ← 892 | `(@[Lott.Symbol_parser, $attrIdent:ident] 893 | private 894 | def $varParserIdent : Parser := 895 | leadingNode $(quote <| ← nt.catName) Parser.maxPrec <| 896 | categoryParser $(quote <| ← nt.varCatName) Parser.maxPrec) 897 | 898 | -- Define parser in parent. 899 | if let some parent := parent? then 900 | let parentQualified := parent.getId 901 | let parentParserCatName := symbolPrefix ++ parentQualified 902 | let parentParserAttrName := parentParserCatName.appendAfter "_parser" 903 | let parentParserIdent := 904 | mkIdentFrom parent <| `_root_ ++ parentQualified ++ canonName.appendAfter "_parser" 905 | elabCommand <| ← 906 | `(@[Lott.Symbol_parser, $(mkIdent parentParserAttrName):ident] 907 | private 908 | def $parentParserIdent : Parser := 909 | leadingNode $(quote parentParserCatName) Parser.maxPrec <| 910 | categoryParser $(quote <| ← nt.catName) 0) 911 | 912 | -- Define macros and tex elaborations. 913 | let catName ← nt.catName 914 | let catIdent := mkIdent catName 915 | for prod@{ name, ir, expand?, ids, profileTex, .. } in prods do 916 | let patternArgs ← IR.toPatternArgs ir 917 | 918 | if ← getTerm then 919 | let macroSeqItems ← IR.toMacroSeqItems ir canon.getId ids prod.binders 920 | let rest ← expand?.getDM do 921 | let exprArgs ← IR.toExprArgs ir ids prod.binders 922 | `(return Syntax.mkCApp $(quote <| ns ++ canonName ++ name.getId) #[$exprArgs,*]) 923 | 924 | let macroName := mkIdentFrom name <| canonName ++ name.getId.appendAfter "Impl" 925 | elabCommand <| ← 926 | `(@[macro $symbolEmbedIdent] 927 | private 928 | def $macroName : Macro := fun stx => do 929 | let Lean.Syntax.node _ ``Lott.symbolEmbed #[ 930 | Lean.Syntax.atom _ "[[", 931 | Lean.Syntax.node _ $(quote catName) #[$patternArgs,*], 932 | Lean.Syntax.atom _ "]]" 933 | ] := stx | Macro.throwUnsupported 934 | $macroSeqItems* 935 | $rest:term) 936 | 937 | if ← getTexOutputSome then 938 | let texSeqItems ← IR.toTexSeqItems ir canon.getId 939 | let profileAlternatives ← profileTex.toArray.filterMapM fun (profile, tex) => do 940 | if profile == default then 941 | return none 942 | return some <| ← `(doSeqItem| if profile == $(quote profile) then return $tex) 943 | let rest ← profileTex.find? default |>.getDM do 944 | let joinArgs := IR.toJoinArgs ir 945 | `(" ".intercalate [$joinArgs,*]) 946 | 947 | let texElabName := mkIdentFrom name <| canonName ++ name.getId.appendAfter "TexElab" 948 | elabCommand <| ← 949 | `(@[lott_tex_elab $catIdent] 950 | private 951 | def $texElabName : TexElab := fun profile ref stx => do 952 | let Lean.Syntax.node _ $(quote catName) #[$patternArgs,*] := stx 953 | | throwUnsupportedSyntax 954 | $texSeqItems* 955 | $profileAlternatives* 956 | return $rest) 957 | 958 | if ← getTerm then 959 | if let some parent := parent? then 960 | let qualified ← nt.qualified 961 | let parentCatName := symbolPrefix ++ parent.getId 962 | let parentCatIdent := mkIdent parentCatName 963 | let parentMacroName := 964 | mkIdent <| `_root_ ++ parent.getId ++ canon.getId.appendAfter "Impl" 965 | elabCommand <| ← 966 | `(@[macro $symbolEmbedIdent] 967 | private 968 | def $parentMacroName : Macro := fun stx => do 969 | let Lean.Syntax.node _ ``Lott.symbolEmbed #[ 970 | Lean.Syntax.atom _ "[[", 971 | Lean.Syntax.node _ $(quote parentCatName) #[ 972 | stx@(Lean.Syntax.node _ $(quote catName) _) 973 | ], 974 | Lean.Syntax.atom _ "]]" 975 | ] := stx | Macro.throwUnsupported 976 | return Lean.Syntax.mkCApp ``Subtype.val #[ 977 | Lean.Syntax.mkTypeAscription (Lott.Syntax.mkSymbolEmbed stx) <| 978 | Lean.mkIdent $(quote qualified) 979 | ]) 980 | 981 | -- Add metavar functions. 982 | if ← getTerm then 983 | let allSubstitutions := nts.map (·.substitutions?.getD #[]) 984 | let uniqueSubstitutions := 985 | allSubstitutions.flatten.foldl Std.HashSet.insert Std.HashSet.empty |>.toArray 986 | for subst@(varTypeName, valTypeName) in uniqueSubstitutions do 987 | let locallyNameless ← isLocallyNameless varTypeName 988 | 989 | let varTypeId := mkIdent <| if locallyNameless then 990 | varTypeName.appendAfter "Id" 991 | else 992 | varTypeName 993 | let valTypeId := mkIdent valTypeName 994 | 995 | let varId ← mkFreshIdent varTypeId 996 | let valId ← mkFreshIdent valTypeId 997 | 998 | let substDefs ← nts.filterMapM fun nt@{ canon, prods, substitutions?, .. } => do 999 | if !(substitutions?.getD #[]).contains subst then 1000 | return none 1001 | 1002 | let matchAlts ← prods.filterMapM fun { name, ir, expand?, bindConfig?, .. } => do 1003 | let none := expand? | return none 1004 | let prodId := mkIdent <| canon.getId ++ name.getId 1005 | 1006 | let bindOf? := bindConfig?.map fun { of, .. } => of 1007 | let binderId? ← bindConfig?.bindM fun bindConfig@{ of, .. } => 1008 | ir.findSomeM? fun (mk l ir) => do 1009 | match ir with 1010 | | .category n => return if n == varTypeName && l == of then some l else none 1011 | | .atom _ => 1012 | if let some ref := bindConfig.find? l.getId then 1013 | throwErrorAt ref "atoms shouldn't be referenced by binders" 1014 | 1015 | return none 1016 | | .sepBy .. 1017 | | .optional _ => 1018 | throwError "bind configuration for productions with sepBy or optional syntax are not supported" 1019 | 1020 | if let #[mk l (.category n)] := ir then 1021 | if some l != binderId? && n == varTypeName then 1022 | let lhs ← if locallyNameless then `(.free $varId) else pure varId 1023 | return ← `(matchAltExpr| | $prodId $l => if $lhs = $l then $valId else $prodId $l) 1024 | 1025 | let some (patternArgs, rhsArgs) ← 1026 | toSubstPatternAndArgs ir varTypeName valTypeName varId valId bindOf? binderId? | 1027 | `(matchAltExpr| | x@($prodId ..) => x) 1028 | 1029 | `(matchAltExpr| | $prodId $patternArgs* => $prodId $rhsArgs*) 1030 | 1031 | return some <| ← 1032 | `(def $(mkIdentFrom canon <| canon.getId ++ (← substName varTypeName)) 1033 | (x : $canon) ($varId : $varTypeId) ($valId : $valTypeId) : $canon := match x with 1034 | $matchAlts:matchAlt*) 1035 | 1036 | elabMutualCommands substDefs 1037 | 1038 | if !locallyNameless then 1039 | continue 1040 | 1041 | let idxId ← mkFreshIdent varTypeId 1042 | 1043 | let varOpenDefs ← nts.filterMapM fun nt@{ canon, prods, substitutions?, .. } => do 1044 | if !(substitutions?.getD #[]).contains subst then 1045 | return none 1046 | 1047 | let matchAlts ← prods.filterMapM fun { name, ir, expand?, bindConfig?, .. } => do 1048 | let none := expand? | return none 1049 | let prodId := mkIdent <| canon.getId ++ name.getId 1050 | 1051 | let bindOf? := bindConfig?.map fun { of, .. } => of 1052 | let binderId? ← bindConfig?.bindM fun bindConfig@{ of, .. } => 1053 | ir.findSomeM? fun (mk l ir) => do 1054 | match ir with 1055 | | .category n => return if n == varTypeName && l == of then some l else none 1056 | | .atom _ => 1057 | if let some ref := bindConfig.find? l.getId then 1058 | throwErrorAt ref "atoms shouldn't be referenced by binders" 1059 | 1060 | return none 1061 | | .sepBy .. 1062 | | .optional _ => 1063 | throwError "bind configuration for productions with sepBy or optional syntax are not supported" 1064 | 1065 | let some (patternArgs, rhsArgs) ← 1066 | toVarOpenPatternAndArgs ir varTypeName valTypeName varId idxId bindOf? binderId? | 1067 | `(matchAltExpr| | x@($prodId ..) => x) 1068 | 1069 | if binderId?.isSome then 1070 | `(matchAltExpr| 1071 | | x@($prodId $patternArgs*) => let $idxId := $idxId + 1; $prodId $rhsArgs*) 1072 | else 1073 | `(matchAltExpr| | $prodId $patternArgs* => $prodId $rhsArgs*) 1074 | 1075 | return some <| ← 1076 | `(def $(mkIdentFrom canon <| canon.getId ++ (← openName varTypeName)) 1077 | (x : $canon) ($varId : $varTypeId) ($idxId : Nat := 0) : $canon := match x with 1078 | $matchAlts:matchAlt*) 1079 | 1080 | elabMutualCommands varOpenDefs 1081 | 1082 | let valOpenDefs ← nts.filterMapM fun nt@{ canon, prods, substitutions?, .. } => do 1083 | if !(substitutions?.getD #[]).contains subst then 1084 | return none 1085 | 1086 | let matchAlts ← prods.filterMapM fun { name, ir, expand?, bindConfig?, .. } => do 1087 | let none := expand? | return none 1088 | let prodId := mkIdent <| canon.getId ++ name.getId 1089 | 1090 | let bindOf? := bindConfig?.map fun { of, .. } => of 1091 | let binderId? ← bindConfig?.bindM fun bindConfig@{ of, .. } => 1092 | ir.findSomeM? fun (mk l ir) => do 1093 | match ir with 1094 | | .category n => return if n == varTypeName && l == of then some l else none 1095 | | .atom _ => 1096 | if let some ref := bindConfig.find? l.getId then 1097 | throwErrorAt ref "atoms shouldn't be referenced by binders" 1098 | 1099 | return none 1100 | | .sepBy .. 1101 | | .optional _ => 1102 | throwError "bind configuration for productions with sepBy or optional syntax are not supported" 1103 | 1104 | if let #[mk l (.category n)] := ir then 1105 | if some l != binderId? && n == varTypeName then 1106 | return ← `(matchAltExpr| | $prodId $l => if .bound $idxId = $l then $valId else $prodId $l) 1107 | 1108 | let some (patternArgs, rhsArgs) ← 1109 | toValOpenPatternAndArgs ir varTypeName valTypeName valId idxId bindOf? binderId? | 1110 | `(matchAltExpr| | x@($prodId ..) => x) 1111 | 1112 | if binderId?.isSome then 1113 | `(matchAltExpr| 1114 | | x@($prodId $patternArgs*) => let $idxId := $idxId + 1; $prodId $rhsArgs*) 1115 | else 1116 | `(matchAltExpr| | $prodId $patternArgs* => $prodId $rhsArgs*) 1117 | 1118 | return some <| ← 1119 | `(def $(mkIdentFrom canon <| canon.getId ++ (← openName valTypeName)) 1120 | (x : $canon) ($valId : $valTypeId) ($idxId : Nat := 0) : $canon := match x with 1121 | $matchAlts:matchAlt*) 1122 | 1123 | elabMutualCommands valOpenDefs 1124 | 1125 | let closeDefs ← nts.filterMapM fun nt@{ canon, prods, substitutions?, .. } => do 1126 | if !(substitutions?.getD #[]).contains subst then 1127 | return none 1128 | 1129 | let matchAlts ← prods.filterMapM fun { name, ir, expand?, bindConfig?, .. } => do 1130 | let none := expand? | return none 1131 | let prodId := mkIdent <| canon.getId ++ name.getId 1132 | 1133 | let bindOf? := bindConfig?.map fun { of, .. } => of 1134 | let binderId? ← bindConfig?.bindM fun bindConfig@{ of, .. } => 1135 | ir.findSomeM? fun (mk l ir) => do 1136 | match ir with 1137 | | .category n => return if n == varTypeName && l == of then some l else none 1138 | | .atom _ => 1139 | if let some ref := bindConfig.find? l.getId then 1140 | throwErrorAt ref "atoms shouldn't be referenced by binders" 1141 | 1142 | return none 1143 | | .sepBy .. 1144 | | .optional _ => 1145 | throwError "bind configuration for productions with sepBy or optional syntax are not supported" 1146 | 1147 | let some (patternArgs, rhsArgs) ← 1148 | toClosePatternAndArgs ir varTypeName valTypeName varId idxId bindOf? binderId? | 1149 | `(matchAltExpr| | x@($prodId ..) => x) 1150 | 1151 | if binderId?.isSome then 1152 | `(matchAltExpr| 1153 | | x@($prodId $patternArgs*) => let $idxId := $idxId + 1; $prodId $rhsArgs*) 1154 | else 1155 | `(matchAltExpr| | $prodId $patternArgs* => $prodId $rhsArgs*) 1156 | 1157 | return some <| ← 1158 | `(def $(mkIdentFrom canon <| canon.getId ++ (← closeName varTypeName)) 1159 | (x : $canon) ($varId : $varTypeId) ($idxId : Nat := 0) : $canon := match x with 1160 | $matchAlts:matchAlt*) 1161 | 1162 | elabMutualCommands closeDefs 1163 | 1164 | let freeDefs ← 1165 | nts.filterMapM fun nt@{ canon, prods, substitutions?, .. } => do 1166 | if !(substitutions?.getD #[]).contains subst then 1167 | return none 1168 | 1169 | let freeId := mkIdent <| ← freeName varTypeName <| ← nt.qualified 1170 | let matchAlts ← prods.filterMapM fun prod@{ name, ir, expand?, .. } => do 1171 | let none := expand? | return none 1172 | let prodId := mkIdentFrom name <| (← nt.qualified) ++ name.getId 1173 | toFreeMatchAlt prodId ir subst prod.binders 1174 | return some <| ← `(def $freeId : $canon → List $varTypeId $matchAlts:matchAlt*) 1175 | 1176 | elabMutualCommands freeDefs 1177 | 1178 | let locallyClosedInductives ← nts.filterMapM fun nt@{ canon, prods, substitutions?, .. } => do 1179 | if !(substitutions?.getD #[]).contains subst then 1180 | return none 1181 | 1182 | let locallyClosedId := mkIdent <| ← locallyClosedName varTypeName <| ← nt.qualified 1183 | let ctors ← prods.mapM fun prod@{ name, ir, expand?, ids, .. } => do 1184 | let none := expand? | return #[] 1185 | toLocallyClosedCtors name locallyClosedId idxId (← nt.qualified) ir subst ids prod.binders 1186 | return some <| ← 1187 | `(inductive $locallyClosedId : $canon → (optParam Nat 0) → Prop where $ctors.flatten*) 1188 | 1189 | elabMutualCommands locallyClosedInductives 1190 | 1191 | -- Write tex output. 1192 | for nt@{ canon, aliases, prods, texPrePost?, profiles, .. } in nts do 1193 | for profile in profiles.push default do 1194 | writeTexOutput (ns ++ canon.getId ++ profile) do 1195 | let canonTex := canon.getId.getFinal.getString!.pascalToTitle.texEscape 1196 | let (texPre, texPost) := texPrePost?.getD ("", "") 1197 | let aliasesTex := "\\lottaliassep\n".intercalate <| aliases.toList.filterMap 1198 | fun (alias, «notex», tex?) => 1199 | if «notex» then 1200 | none 1201 | else 1202 | let aliasTex := tex?.getD alias.getId.getFinal.getString!.texEscape 1203 | s!"{texPre}\\lottalias\{{aliasTex}}{texPost}\n" 1204 | let productionTexs ← prods.filterMapM fun { name, ir, «notex», .. } => do 1205 | if «notex» then return none 1206 | 1207 | let canonQualified := ns ++ canon.getId 1208 | let catName ← nt.catName 1209 | let exampleStx := mkNode catName <| ← toExampleSyntax ir canonQualified profile 1210 | let productionTex ← liftTermElabM <| 1211 | texElabSymbolOrJudgement catName profile name exampleStx 1212 | return some s!"\\lottproduction\{{productionTex}}\n" 1213 | let productionsTex := "\\lottproductionsep\n".intercalate productionTexs.toList 1214 | return s!"\\lottnonterminal\{{canonTex}}\{\n{aliasesTex}}\{\n{productionsTex}}\n" 1215 | 1216 | elab_rules : command 1217 | | `($nt:Lott.NonTerminal) => elabNonTerminals #[nt] 1218 | | `(mutual $[$nts:Lott.NonTerminal]* end) => elabNonTerminals nts 1219 | 1220 | end Lott 1221 | --------------------------------------------------------------------------------