├── .github └── workflows │ └── ci.yml ├── .gitignore ├── Export.lean ├── LICENSE ├── Main.lean ├── README.md ├── Test.lean ├── examples ├── CategoryTheory.yonedaLemma.export └── Nat.gcd_self.export ├── lake-manifest.json ├── lakefile.toml └── lean-toolchain /.github/workflows/ci.yml: -------------------------------------------------------------------------------- 1 | name: CI 2 | 3 | on: 4 | push: 5 | branches: ["master"] 6 | pull_request: 7 | branches: ["master"] 8 | workflow_dispatch: 9 | 10 | jobs: 11 | build: 12 | runs-on: ubuntu-latest 13 | 14 | steps: 15 | - uses: actions/checkout@v4 16 | - uses: leanprover/lean-action@v1 17 | -------------------------------------------------------------------------------- /.gitignore: -------------------------------------------------------------------------------- 1 | .lake/ 2 | -------------------------------------------------------------------------------- /Export.lean: -------------------------------------------------------------------------------- 1 | import Lean 2 | import Std.Data.HashMap.Basic 3 | 4 | open Lean 5 | open Std (HashMap) 6 | 7 | instance instHashableRecursorRule : Hashable RecursorRule where 8 | hash r := hash (r.ctor, r.nfields, r.rhs) 9 | 10 | structure Context where 11 | env : Environment 12 | 13 | structure State where 14 | visitedNames : HashMap Name Nat := .insert {} .anonymous 0 15 | visitedLevels : HashMap Level Nat := .insert {} .zero 0 16 | visitedExprs : HashMap Expr Nat := {} 17 | visitedRecRules : HashMap RecursorRule Nat := {} 18 | visitedConstants : NameHashSet := {} 19 | noMDataExprs : HashMap Expr Expr := {} 20 | exportUnsafe : Bool := false 21 | 22 | abbrev M := ReaderT Context <| StateT State IO 23 | 24 | def M.run (env : Environment) (act : M α) : IO α := 25 | StateT.run' (s := {}) do 26 | ReaderT.run (r := { env }) do 27 | act 28 | 29 | @[inline] 30 | def getIdx [Hashable α] [BEq α] (x : α) (getM : State → HashMap α Nat) (setM : State → HashMap α Nat → State) (rec : M String) : M Nat := do 31 | let m ← getM <$> get 32 | if let some idx := m[x]? then 33 | return idx 34 | let s ← rec 35 | let m ← getM <$> get 36 | let idx := m.size 37 | IO.println s!"{idx} {s}" 38 | modify fun st => setM st ((getM st).insert x idx) 39 | return idx 40 | 41 | def dumpName (n : Name) : M Nat := getIdx n (·.visitedNames) ({ · with visitedNames := · }) do 42 | match n with 43 | | .anonymous => unreachable! 44 | | .str n s => return s!"#NS {← dumpName n} {s}" 45 | | .num n i => return s!"#NI {← dumpName n} {i}" 46 | 47 | def dumpLevel (l : Level) : M Nat := getIdx l (·.visitedLevels) ({ · with visitedLevels := · }) do 48 | match l with 49 | | .zero | .mvar _ => unreachable! 50 | | .succ l => return s!"#US {← dumpLevel l}" 51 | | .max l1 l2 => return s!"#UM {← dumpLevel l1} {← dumpLevel l2}" 52 | | .imax l1 l2 => return s!"#UIM {← dumpLevel l1} {← dumpLevel l2}" 53 | | .param n => return s!"#UP {← dumpName n}" 54 | 55 | def seq [ToString α] (xs : List α) : String := 56 | xs.map toString |> String.intercalate " " 57 | 58 | def dumpInfo : BinderInfo → String 59 | | .default => "#BD" 60 | | .implicit => "#BI" 61 | | .strictImplicit => "#BS" 62 | | .instImplicit => "#BC" 63 | 64 | def uint8ToHex (c : UInt8) : String := 65 | let d2 := c / 16 66 | let d1 := c % 16 67 | (hexDigitRepr d2.toNat ++ hexDigitRepr d1.toNat).toUpper 68 | 69 | def removeMData (e : Expr) : M Expr := do 70 | if let some x := (← get).noMDataExprs[e]? then 71 | return x 72 | let e' ← match e with 73 | | .mdata _ e' => removeMData e' 74 | | .fvar .. | .mvar ..=> unreachable! 75 | | .bvar .. | .sort .. | .const .. | .lit _ => pure e 76 | | .app f a => 77 | pure <| e.updateApp! (← removeMData f) (← removeMData a) 78 | | .lam _ d b _ => 79 | pure <| e.updateLambdaE! (← removeMData d) (← removeMData b) 80 | | .letE _ d v b _ => 81 | pure <| e.updateLet! (← removeMData d) (← removeMData v) (← removeMData b) 82 | | .forallE _ d b _ => 83 | pure <| e.updateForallE! (← removeMData d) (← removeMData b) 84 | | .proj _ _ e2 => 85 | pure <| e.updateProj! (← removeMData e2) 86 | modify (fun st => { st with noMDataExprs := st.noMDataExprs.insert e e' }) 87 | pure e' 88 | 89 | partial def dumpExprAux (e : Expr) : M Nat := do 90 | getIdx e (·.visitedExprs) ({ · with visitedExprs := · }) do 91 | match e with 92 | | .bvar i => return s!"#EV {i}" 93 | | .sort l => return s!"#ES {← dumpLevel l}" 94 | | .const n us => 95 | return s!"#EC {← dumpName n} {← seq <$> us.mapM dumpLevel}" 96 | | .app f e => return s!"#EA {← dumpExprAux f} {← dumpExprAux e}" 97 | | .lam n d b bi => return s!"#EL {dumpInfo bi} {← dumpName n} {← dumpExprAux d} {← dumpExprAux b}" 98 | | .letE n d v b _ => return s!"#EZ {← dumpName n} {← dumpExprAux d} {← dumpExprAux v} {← dumpExprAux b}" 99 | | .forallE n d b bi => return s!"#EP {dumpInfo bi} {← dumpName n} {← dumpExprAux d} {← dumpExprAux b}" 100 | | .mdata .. | .fvar .. | .mvar .. => unreachable! 101 | -- extensions compared to Lean 3 102 | | .proj s i e => return s!"#EJ {← dumpName s} {i} {← dumpExprAux e}" 103 | | .lit (.natVal i) => return s!"#ELN {i}" 104 | | .lit (.strVal s) => return s!"#ELS {s.toUTF8.toList.map uint8ToHex |> seq}" 105 | 106 | def dumpExpr (e : Expr) : M Nat := do 107 | dumpExprAux (← removeMData e) 108 | 109 | def dumpHints : ReducibilityHints → String 110 | | .opaque => "O" 111 | | .abbrev => "A" 112 | | .regular n => s!"R {n}" 113 | 114 | def dumpUparams (uparams : List Name) : M (List Nat) := do 115 | let nameIdxs ← uparams.mapM dumpName 116 | let _ ← (uparams.map (.param)).mapM dumpLevel 117 | pure nameIdxs 118 | 119 | partial def dumpConstant (c : Name) : M Unit := do 120 | let declar := ((← read).env.find? c).get! 121 | if (declar.isUnsafe && !(← get).exportUnsafe) || (← get).visitedConstants.contains c then 122 | return 123 | modify fun st => { st with visitedConstants := st.visitedConstants.insert c } 124 | match declar with 125 | | .axiomInfo val => do 126 | IO.println s!"#AX {← dumpName c} {← dumpExpr val.type} {← seq <$> dumpUparams val.levelParams}" 127 | | .defnInfo val => do 128 | dumpDeps val.type 129 | dumpDeps val.value 130 | IO.println s!"#DEF {← dumpName c} {← dumpExpr val.type} {← dumpExpr val.value} {dumpHints val.hints} {← seq <$> dumpUparams val.levelParams}" 131 | | .thmInfo val => do 132 | dumpDeps val.type 133 | dumpDeps val.value 134 | IO.println s!"#THM {← dumpName c} {← dumpExpr val.type} {← dumpExpr val.value} {← seq <$> dumpUparams val.levelParams}" 135 | | .opaqueInfo val => do 136 | dumpDeps val.type 137 | dumpDeps val.value 138 | IO.println s!"#OPAQ {← dumpName c} {← dumpExpr val.type} {← dumpExpr val.value} {← seq <$> dumpUparams val.levelParams}" 139 | | .quotInfo val => 140 | dumpDeps val.type 141 | IO.println s!"#QUOT {← dumpName c} {← dumpExpr val.type} {← seq <$> dumpUparams val.levelParams}" 142 | | .inductInfo val => do 143 | dumpDeps val.type 144 | let indNameIdxs ← val.all.mapM dumpName 145 | let ctorNameIdxs ← val.ctors.mapM (fun ctor => dumpName ctor) 146 | let isReflexive := if val.isReflexive then 1 else 0 147 | let isRec := if val.isRec then 1 else 0 148 | IO.println s!"#IND {← dumpName c} {← dumpExpr val.type} {isReflexive} {isRec} {val.numNested} {val.numParams} {val.numIndices} {indNameIdxs.length} {seq indNameIdxs} {val.numCtors} {seq ctorNameIdxs} {← seq <$> dumpUparams val.levelParams}" 149 | /- 150 | We now have at least one inductive exported for which the constructor is never invoked, but 151 | they're needed for the recursors. 152 | -/ 153 | for ctor in val.ctors do 154 | dumpConstant ctor 155 | | .ctorInfo val => 156 | dumpDeps val.type 157 | IO.println s!"#CTOR {← dumpName c} {← dumpExpr val.type} {← dumpName val.induct} {val.cidx} {val.numParams} {val.numFields} {← seq <$> dumpUparams val.levelParams}" 158 | | .recInfo val => 159 | dumpDeps val.type 160 | let indNameIdxs ← val.all.mapM dumpName 161 | let k := if val.k then 1 else 0 162 | let ruleIdxs ← val.rules.mapM (fun rule => dumpRecRule rule) 163 | IO.println s!"#REC {← dumpName c} {← dumpExpr val.type} {indNameIdxs.length} {seq indNameIdxs} {val.numParams} {val.numIndices} {val.numMotives} {val.numMinors} {ruleIdxs.length} {seq ruleIdxs} {k} {← seq <$> dumpUparams val.levelParams}" 164 | where 165 | dumpDeps e := do 166 | for c in e.getUsedConstants do 167 | dumpConstant c 168 | dumpRecRule (rule : RecursorRule) : M Nat := getIdx rule (·.visitedRecRules) ({ · with visitedRecRules := · }) do 169 | dumpDeps (rule.rhs) 170 | return s!"#RR {← dumpName rule.ctor} {rule.nfields} {← dumpExpr rule.rhs}" 171 | -------------------------------------------------------------------------------- /LICENSE: -------------------------------------------------------------------------------- 1 | Apache License 2 | Version 2.0, January 2004 3 | http://www.apache.org/licenses/ 4 | 5 | TERMS AND CONDITIONS FOR USE, REPRODUCTION, AND DISTRIBUTION 6 | 7 | 1. Definitions. 8 | 9 | "License" shall mean the terms and conditions for use, reproduction, 10 | and distribution as defined by Sections 1 through 9 of this document. 11 | 12 | "Licensor" shall mean the copyright owner or entity authorized by 13 | the copyright owner that is granting the License. 14 | 15 | "Legal Entity" shall mean the union of the acting entity and all 16 | other entities that control, are controlled by, or are under common 17 | control with that entity. For the purposes of this definition, 18 | "control" means (i) the power, direct or indirect, to cause the 19 | direction or management of such entity, whether by contract or 20 | otherwise, or (ii) ownership of fifty percent (50%) or more of the 21 | outstanding shares, or (iii) beneficial ownership of such entity. 22 | 23 | "You" (or "Your") shall mean an individual or Legal Entity 24 | exercising permissions granted by this License. 25 | 26 | "Source" form shall mean the preferred form for making modifications, 27 | including but not limited to software source code, documentation 28 | source, and configuration files. 29 | 30 | "Object" form shall mean any form resulting from mechanical 31 | transformation or translation of a Source form, including but 32 | not limited to compiled object code, generated documentation, 33 | and conversions to other media types. 34 | 35 | "Work" shall mean the work of authorship, whether in Source or 36 | Object form, made available under the License, as indicated by a 37 | copyright notice that is included in or attached to the work 38 | (an example is provided in the Appendix below). 39 | 40 | "Derivative Works" shall mean any work, whether in Source or Object 41 | form, that is based on (or derived from) the Work and for which the 42 | editorial revisions, annotations, elaborations, or other modifications 43 | represent, as a whole, an original work of authorship. For the purposes 44 | of this License, Derivative Works shall not include works that remain 45 | separable from, or merely link (or bind by name) to the interfaces of, 46 | the Work and Derivative Works thereof. 47 | 48 | "Contribution" shall mean any work of authorship, including 49 | the original version of the Work and any modifications or additions 50 | to that Work or Derivative Works thereof, that is intentionally 51 | submitted to Licensor for inclusion in the Work by the copyright owner 52 | or by an individual or Legal Entity authorized to submit on behalf of 53 | the copyright owner. For the purposes of this definition, "submitted" 54 | means any form of electronic, verbal, or written communication sent 55 | to the Licensor or its representatives, including but not limited to 56 | communication on electronic mailing lists, source code control systems, 57 | and issue tracking systems that are managed by, or on behalf of, the 58 | Licensor for the purpose of discussing and improving the Work, but 59 | excluding communication that is conspicuously marked or otherwise 60 | designated in writing by the copyright owner as "Not a Contribution." 61 | 62 | "Contributor" shall mean Licensor and any individual or Legal Entity 63 | on behalf of whom a Contribution has been received by Licensor and 64 | subsequently incorporated within the Work. 65 | 66 | 2. Grant of Copyright License. Subject to the terms and conditions of 67 | this License, each Contributor hereby grants to You a perpetual, 68 | worldwide, non-exclusive, no-charge, royalty-free, irrevocable 69 | copyright license to reproduce, prepare Derivative Works of, 70 | publicly display, publicly perform, sublicense, and distribute the 71 | Work and such Derivative Works in Source or Object form. 72 | 73 | 3. Grant of Patent License. Subject to the terms and conditions of 74 | this License, each Contributor hereby grants to You a perpetual, 75 | worldwide, non-exclusive, no-charge, royalty-free, irrevocable 76 | (except as stated in this section) patent license to make, have made, 77 | use, offer to sell, sell, import, and otherwise transfer the Work, 78 | where such license applies only to those patent claims licensable 79 | by such Contributor that are necessarily infringed by their 80 | Contribution(s) alone or by combination of their Contribution(s) 81 | with the Work to which such Contribution(s) was submitted. If You 82 | institute patent litigation against any entity (including a 83 | cross-claim or counterclaim in a lawsuit) alleging that the Work 84 | or a Contribution incorporated within the Work constitutes direct 85 | or contributory patent infringement, then any patent licenses 86 | granted to You under this License for that Work shall terminate 87 | as of the date such litigation is filed. 88 | 89 | 4. Redistribution. You may reproduce and distribute copies of the 90 | Work or Derivative Works thereof in any medium, with or without 91 | modifications, and in Source or Object form, provided that You 92 | meet the following conditions: 93 | 94 | (a) You must give any other recipients of the Work or 95 | Derivative Works a copy of this License; and 96 | 97 | (b) You must cause any modified files to carry prominent notices 98 | stating that You changed the files; and 99 | 100 | (c) You must retain, in the Source form of any Derivative Works 101 | that You distribute, all copyright, patent, trademark, and 102 | attribution notices from the Source form of the Work, 103 | excluding those notices that do not pertain to any part of 104 | the Derivative Works; and 105 | 106 | (d) If the Work includes a "NOTICE" text file as part of its 107 | distribution, then any Derivative Works that You distribute must 108 | include a readable copy of the attribution notices contained 109 | within such NOTICE file, excluding those notices that do not 110 | pertain to any part of the Derivative Works, in at least one 111 | of the following places: within a NOTICE text file distributed 112 | as part of the Derivative Works; within the Source form or 113 | documentation, if provided along with the Derivative Works; or, 114 | within a display generated by the Derivative Works, if and 115 | wherever such third-party notices normally appear. The contents 116 | of the NOTICE file are for informational purposes only and 117 | do not modify the License. You may add Your own attribution 118 | notices within Derivative Works that You distribute, alongside 119 | or as an addendum to the NOTICE text from the Work, provided 120 | that such additional attribution notices cannot be construed 121 | as modifying the License. 122 | 123 | You may add Your own copyright statement to Your modifications and 124 | may provide additional or different license terms and conditions 125 | for use, reproduction, or distribution of Your modifications, or 126 | for any such Derivative Works as a whole, provided Your use, 127 | reproduction, and distribution of the Work otherwise complies with 128 | the conditions stated in this License. 129 | 130 | 5. Submission of Contributions. Unless You explicitly state otherwise, 131 | any Contribution intentionally submitted for inclusion in the Work 132 | by You to the Licensor shall be under the terms and conditions of 133 | this License, without any additional terms or conditions. 134 | Notwithstanding the above, nothing herein shall supersede or modify 135 | the terms of any separate license agreement you may have executed 136 | with Licensor regarding such Contributions. 137 | 138 | 6. Trademarks. This License does not grant permission to use the trade 139 | names, trademarks, service marks, or product names of the Licensor, 140 | except as required for reasonable and customary use in describing the 141 | origin of the Work and reproducing the content of the NOTICE file. 142 | 143 | 7. Disclaimer of Warranty. Unless required by applicable law or 144 | agreed to in writing, Licensor provides the Work (and each 145 | Contributor provides its Contributions) on an "AS IS" BASIS, 146 | WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or 147 | implied, including, without limitation, any warranties or conditions 148 | of TITLE, NON-INFRINGEMENT, MERCHANTABILITY, or FITNESS FOR A 149 | PARTICULAR PURPOSE. You are solely responsible for determining the 150 | appropriateness of using or redistributing the Work and assume any 151 | risks associated with Your exercise of permissions under this License. 152 | 153 | 8. Limitation of Liability. In no event and under no legal theory, 154 | whether in tort (including negligence), contract, or otherwise, 155 | unless required by applicable law (such as deliberate and grossly 156 | negligent acts) or agreed to in writing, shall any Contributor be 157 | liable to You for damages, including any direct, indirect, special, 158 | incidental, or consequential damages of any character arising as a 159 | result of this License or out of the use or inability to use the 160 | Work (including but not limited to damages for loss of goodwill, 161 | work stoppage, computer failure or malfunction, or any and all 162 | other commercial damages or losses), even if such Contributor 163 | has been advised of the possibility of such damages. 164 | 165 | 9. Accepting Warranty or Additional Liability. While redistributing 166 | the Work or Derivative Works thereof, You may choose to offer, 167 | and charge a fee for, acceptance of support, warranty, indemnity, 168 | or other liability obligations and/or rights consistent with this 169 | License. However, in accepting such obligations, You may act only 170 | on Your own behalf and on Your sole responsibility, not on behalf 171 | of any other Contributor, and only if You agree to indemnify, 172 | defend, and hold each Contributor harmless for any liability 173 | incurred by, or claims asserted against, such Contributor by reason 174 | of your accepting any such warranty or additional liability. 175 | 176 | END OF TERMS AND CONDITIONS 177 | 178 | APPENDIX: How to apply the Apache License to your work. 179 | 180 | To apply the Apache License to your work, attach the following 181 | boilerplate notice, with the fields enclosed by brackets "[]" 182 | replaced with your own identifying information. (Don't include 183 | the brackets!) The text should be enclosed in the appropriate 184 | comment syntax for the file format. We also recommend that a 185 | file or class name and description of purpose be included on the 186 | same "printed page" as the copyright notice for easier 187 | identification within third-party archives. 188 | 189 | Copyright [yyyy] [name of copyright owner] 190 | 191 | Licensed under the Apache License, Version 2.0 (the "License"); 192 | you may not use this file except in compliance with the License. 193 | You may obtain a copy of the License at 194 | 195 | http://www.apache.org/licenses/LICENSE-2.0 196 | 197 | Unless required by applicable law or agreed to in writing, software 198 | distributed under the License is distributed on an "AS IS" BASIS, 199 | WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. 200 | See the License for the specific language governing permissions and 201 | limitations under the License. 202 | -------------------------------------------------------------------------------- /Main.lean: -------------------------------------------------------------------------------- 1 | import Export 2 | open Lean 3 | 4 | def semver := "2.0.0" 5 | 6 | def main (args : List String) : IO Unit := do 7 | initSearchPath (← findSysroot) 8 | let (opts, args) := args.partition (fun s => s.startsWith "--" && s.length ≥ 3) 9 | let (imports, constants) := args.span (· != "--") 10 | let imports := imports.toArray.map fun mod => { module := Syntax.decodeNameLit ("`" ++ mod) |>.get! } 11 | let env ← importModules imports {} 12 | let constants := match constants.tail? with 13 | | some cs => cs.map fun c => Syntax.decodeNameLit ("`" ++ c) |>.get! 14 | | none => env.constants.toList.map Prod.fst |>.filter (!·.isInternal) 15 | M.run env do 16 | modify (fun st => { st with exportUnsafe := opts.any (· == "--export-unsafe") }) 17 | IO.println semver 18 | for c in constants do 19 | modify (fun st => { st with noMDataExprs := {} }) 20 | let _ ← dumpConstant c 21 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | A simple declaration exporter for Lean 4 using the [Lean 4 export format](https://ammkrn.github.io/type_checking_in_lean4/export_format.html) 2 | 3 | ## How to Run 4 | 5 | ```sh 6 | $ lake exe lean4export [options] [-- ] 7 | ``` 8 | This exports the contents of the given Lean modules, looked up in the core library or `LEAN_PATH` (as e.g. initialized by an outer `lake env`) and their transitive dependencies. 9 | A specific list of declarations to be exported from these modules can be given after a separating `--`. 10 | The option `--export-unsafe` can be used to include unsafe declarations in the export file. This may be useful for testing and debugging other tools, where unsafe declarations can serve as negative examples. 11 | 12 | 13 | ## Format Extensions 14 | 15 | The following commands have been added to represent new features of the Lean 4 type system. 16 | 17 | ``` 18 | #EJ 19 | ``` 20 | A primitive projection of the ``-nth field of a value `` of the record type ``. 21 | Example: the primitive projection corresponding to `Prod.snd` of the innermost bound variable 22 | ``` 23 | 1 #NS 0 Prod 24 | 0 #EV 0 25 | 1 #EJ 1 1 0 26 | ``` 27 | 28 | ``` 29 | #ELN 30 | #ELS * 31 | ``` 32 | Primitive literals of type `Nat` and `String` (encoded as a sequence of UTF-8 bytes in hexadecimal), respectively. 33 | Examples: 34 | ``` 35 | 0 #ELN 1000000000000000 36 | 1 #ELS 68 69 # "hi" 37 | ``` 38 | 39 | ``` 40 | #EZ 41 | ``` 42 | A binding `let : := ; `. 43 | Already supported by the Lean 3 export format, but not documented. 44 | Example: the encoding of `let x : Nat := Nat.zero; x` is 45 | ``` 46 | 1 #NS 0 x 47 | 2 #NS 0 Nat 48 | 0 #EC 2 49 | 3 #NS 2 zero 50 | 1 #EC 3 51 | 2 #EV 0 52 | 3 #EZ 1 0 1 2 53 | ``` 54 | -------------------------------------------------------------------------------- /Test.lean: -------------------------------------------------------------------------------- 1 | import Export 2 | open Lean 3 | 4 | def run (act : M α) : MetaM Unit := do let _ ← M.run (← getEnv) act 5 | 6 | /-- 7 | info: 1 #NS 0 foo 8 | 2 #NS 1 bla 9 | 3 #NI 2 1 10 | 4 #NS 3 boo 11 | -/ 12 | #guard_msgs in 13 | #eval run <| dumpName (`foo.bla |>.num 1 |>.str "boo") 14 | 15 | /-- 16 | info: 1 #US 0 17 | 2 #US 1 18 | 1 #NS 0 l1 19 | 3 #UP 1 20 | 4 #UM 2 3 21 | 2 #NS 0 l2 22 | 5 #UP 2 23 | 6 #UIM 4 5 24 | -/ 25 | #guard_msgs in 26 | #eval run <| dumpLevel (.imax (.max (.succ (.succ .zero)) (.param `l1)) (.param `l2)) 27 | 28 | /-- 29 | info: 1 #NS 0 A 30 | 1 #US 0 31 | 0 #ES 1 32 | 2 #NS 0 a 33 | 1 #EV 0 34 | 2 #EL #BD 2 1 1 35 | 3 #EL #BI 1 0 2 36 | -/ 37 | #guard_msgs in 38 | #eval run <| dumpExpr (.lam `A (.sort (.succ .zero)) (.lam `a (.bvar 0) (.bvar 0) .default) .implicit) 39 | 40 | /-- 41 | info: 1 #NS 0 x 42 | 2 #NS 0 Nat 43 | 0 #EC 2 ⏎ 44 | 3 #NS 2 zero 45 | 1 #EC 3 ⏎ 46 | 2 #EV 0 47 | 3 #EZ 1 0 1 2 48 | -/ 49 | #guard_msgs in 50 | #eval run <| dumpExpr (.letE `x (.const `Nat []) (.const `Nat.zero []) (.bvar 0) false) 51 | 52 | /-- 53 | info: 1 #NS 0 Prod 54 | 0 #EV 0 55 | 1 #EJ 1 1 0 56 | -/ 57 | #guard_msgs in 58 | #eval run <| dumpExpr (.proj `Prod 1 (.bvar 0)) 59 | 60 | /-- info: 0 #ELN 1000000000000000 -/ 61 | #guard_msgs in 62 | #eval run <| dumpExpr (.lit (.natVal 1000000000000000)) 63 | 64 | /-- info: 0 #ELS 68 69 -/ 65 | #guard_msgs in 66 | #eval run <| dumpExpr (.lit (.strVal "hi")) 67 | 68 | /-- 69 | info: 1 #NS 0 id 70 | 2 #NS 0 α 71 | 3 #NS 0 u 72 | 1 #UP 3 73 | 0 #ES 1 74 | 4 #NS 0 a 75 | 1 #EV 0 76 | 2 #EV 1 77 | 3 #EP #BD 4 1 2 78 | 4 #EP #BI 2 0 3 79 | 5 #EL #BD 4 1 1 80 | 6 #EL #BI 2 0 5 81 | #DEF 1 4 6 R 1 3 82 | -/ 83 | #guard_msgs in 84 | #eval run <| dumpConstant `id 85 | 86 | /-- 87 | info: 1 #NS 0 List 88 | 2 #NS 1 nil 89 | 3 #NS 1 cons 90 | 4 #NS 0 α 91 | 5 #NS 0 u 92 | 1 #UP 5 93 | 2 #US 1 94 | 0 #ES 2 95 | 1 #EP #BD 4 0 0 96 | #IND 1 1 0 1 0 1 0 1 1 2 2 3 5 97 | 2 #EC 1 1 98 | 3 #EV 0 99 | 4 #EA 2 3 100 | 5 #EP #BI 4 0 4 101 | #CTOR 2 5 1 0 1 0 5 102 | 6 #NS 0 head 103 | 7 #NS 0 tail 104 | 6 #EV 1 105 | 7 #EA 2 6 106 | 8 #EV 2 107 | 9 #EA 2 8 108 | 10 #EP #BD 7 7 9 109 | 11 #EP #BD 6 3 10 110 | 12 #EP #BI 4 0 11 111 | #CTOR 3 12 1 1 1 2 5 112 | -/ 113 | #guard_msgs in 114 | #eval run <| dumpConstant `List 115 | -------------------------------------------------------------------------------- /lake-manifest.json: -------------------------------------------------------------------------------- 1 | {"version": "1.1.0", 2 | "packagesDir": ".lake/packages", 3 | "packages": [], 4 | "name": "lean4export", 5 | "lakeDir": ".lake"} 6 | -------------------------------------------------------------------------------- /lakefile.toml: -------------------------------------------------------------------------------- 1 | name = "lean4export" 2 | defaultTargets = ["lean4export"] 3 | testDriver = "Test" 4 | 5 | [[lean_lib]] 6 | name = "Export" 7 | 8 | [[lean_lib]] 9 | name = "Test" 10 | 11 | [[lean_exe]] 12 | name = "lean4export" 13 | root = "Main" 14 | supportInterpreter = true 15 | -------------------------------------------------------------------------------- /lean-toolchain: -------------------------------------------------------------------------------- 1 | leanprover/lean4:v4.21.0-rc3 2 | --------------------------------------------------------------------------------