style: refactor process to Lean 4 compiler style

main
Henrik Böving 2023-01-01 19:30:28 +01:00
parent a5caefc03f
commit 3a5c0db46b
14 changed files with 89 additions and 92 deletions

View File

@ -96,7 +96,7 @@ def getAllModuleDocs (relevantModules : Array Name) : MetaM (HashMap Name Module
let moduleData := env.header.moduleData.get! modIdx let moduleData := env.header.moduleData.get! modIdx
let imports := moduleData.imports.map Import.module let imports := moduleData.imports.map Import.module
res := res.insert module <| Module.mk module modDocs imports res := res.insert module <| Module.mk module modDocs imports
pure res return res
/-- /--
Run the doc-gen analysis on all modules that are loaded into the `Environment` Run the doc-gen analysis on all modules that are loaded into the `Environment`
@ -120,17 +120,17 @@ def process (task : AnalyzeTask) : MetaM (AnalyzerResult × Hierarchy) := do
try try
let config := { let config := {
maxHeartbeats := 5000000, maxHeartbeats := 5000000,
options := ←getOptions, options := ← getOptions,
fileName := ←getFileName, fileName := ← getFileName,
fileMap := ←getFileMap fileMap := ← getFileMap
} }
let analysis := Prod.fst <$> Meta.MetaM.toIO (DocInfo.ofConstant (name, cinfo)) config { env := env } {} {} let analysis Prod.fst <$> Meta.MetaM.toIO (DocInfo.ofConstant (name, cinfo)) config { env := env } {} {}
if let some dinfo analysis then if let some dinfo := analysis then
let moduleName := env.allImportedModuleNames.get! modidx let moduleName := env.allImportedModuleNames.get! modidx
let module := res.find! moduleName let module := res.find! moduleName
res := res.insert moduleName {module with members := module.members.push (ModuleMember.docInfo dinfo)} res := res.insert moduleName {module with members := module.members.push (ModuleMember.docInfo dinfo)}
catch e => catch e =>
IO.println s!"WARNING: Failed to obtain information for: {name}: {←e.toMessageData.toString}" IO.println s!"WARNING: Failed to obtain information for: {name}: {← e.toMessageData.toString}"
-- TODO: This could probably be faster if we did sorted insert above instead -- TODO: This could probably be faster if we did sorted insert above instead
for (moduleName, module) in res.toArray do for (moduleName, module) in res.toArray do
@ -142,7 +142,7 @@ def process (task : AnalyzeTask) : MetaM (AnalyzerResult × Hierarchy) := do
moduleNames := allModules, moduleNames := allModules,
moduleInfo := res, moduleInfo := res,
} }
pure (analysis, hierarchy) return (analysis, hierarchy)
def filterMapDocInfo (ms : Array ModuleMember) : Array DocInfo := def filterMapDocInfo (ms : Array ModuleMember) : Array DocInfo :=
ms.filterMap filter ms.filterMap filter

View File

@ -102,18 +102,18 @@ def parametricAttributes : Array ParametricAttrWrapper := #[⟨externAttr⟩,
def getTags (decl : Name) : MetaM (Array String) := do def getTags (decl : Name) : MetaM (Array String) := do
let env ← getEnv let env ← getEnv
pure <| tagAttributes.filter (TagAttribute.hasTag · env decl) |>.map (λ t => t.attr.name.toString) return tagAttributes.filter (TagAttribute.hasTag · env decl) |>.map (·.attr.name.toString)
def getValuesAux {α : Type} {attrKind : Type → Type} [va : ValueAttr attrKind] [Inhabited α] [ToString α] (decl : Name) (attr : attrKind α) : MetaM (Option String) := do def getValuesAux {α : Type} {attrKind : Type → Type} [va : ValueAttr attrKind] [Inhabited α] [ToString α] (decl : Name) (attr : attrKind α) : MetaM (Option String) := do
let env ← getEnv let env ← getEnv
pure <| va.getValue attr env decl return va.getValue attr env decl
def getValues {attrKind : Type → Type} [ValueAttr attrKind] (decl : Name) (attrs : Array (ValueAttrWrapper attrKind)) : MetaM (Array String) := do def getValues {attrKind : Type → Type} [ValueAttr attrKind] (decl : Name) (attrs : Array (ValueAttrWrapper attrKind)) : MetaM (Array String) := do
let mut res := #[] let mut res := #[]
for attr in attrs do for attr in attrs do
if let some val ← @getValuesAux attr.α attrKind _ attr.inhab attr.str decl attr.attr then if let some val ← @getValuesAux attr.α attrKind _ attr.inhab attr.str decl attr.attr then
res := res.push val res := res.push val
pure res return res
def getEnumValues (decl : Name) : MetaM (Array String) := getValues decl enumAttributes def getEnumValues (decl : Name) : MetaM (Array String) := getValues decl enumAttributes
def getParametricValues (decl : Name) : MetaM (Array String) := getValues decl parametricAttributes def getParametricValues (decl : Name) : MetaM (Array String) := getValues decl parametricAttributes
@ -123,23 +123,21 @@ def getDefaultInstance (decl : Name) (className : Name) : MetaM (Option String)
for (inst, prio) in insts do for (inst, prio) in insts do
if inst == decl then if inst == decl then
return some s!"defaultInstance {prio}" return some s!"defaultInstance {prio}"
pure none return none
def hasSimp (decl : Name) : MetaM (Option String) := do def hasSimp (decl : Name) : MetaM (Option String) := do
let thms ← simpExtension.getTheorems let thms ← simpExtension.getTheorems
pure <| if thms.isLemma (.decl decl) then
if thms.isLemma (.decl decl) then return "simp"
some "simp" else
else return none
none
def hasCsimp (decl : Name) : MetaM (Option String) := do def hasCsimp (decl : Name) : MetaM (Option String) := do
let env ← getEnv let env ← getEnv
pure <| if Compiler.hasCSimpAttribute env decl then
if Compiler.hasCSimpAttribute env decl then return some "csimp"
some "csimp" else
else return none
none
/-- /--
The list of custom attributes, that don't fit in the parametric or enum The list of custom attributes, that don't fit in the parametric or enum
@ -152,7 +150,7 @@ def getCustomAttrs (decl : Name) : MetaM (Array String) := do
for attr in customAttrs do for attr in customAttrs do
if let some value ← attr decl then if let some value ← attr decl then
values := values.push value values := values.push value
pure values return values
/-- /--
The main entry point for recovering all attribute values for a given The main entry point for recovering all attribute values for a given
@ -163,6 +161,6 @@ def getAllAttributes (decl : Name) : MetaM (Array String) := do
let enums ← getEnumValues decl let enums ← getEnumValues decl
let parametric ← getParametricValues decl let parametric ← getParametricValues decl
let customs ← getCustomAttrs decl let customs ← getCustomAttrs decl
pure <| customs ++ tags ++ enums ++ parametric return customs ++ tags ++ enums ++ parametric
end DocGen4 end DocGen4

View File

@ -14,7 +14,7 @@ open Lean Meta
def AxiomInfo.ofAxiomVal (v : AxiomVal) : MetaM AxiomInfo := do def AxiomInfo.ofAxiomVal (v : AxiomVal) : MetaM AxiomInfo := do
let info ← Info.ofConstantVal v.toConstantVal let info ← Info.ofConstantVal v.toConstantVal
pure { return {
toInfo := info, toInfo := info,
isUnsafe := v.isUnsafe isUnsafe := v.isUnsafe
} }

View File

@ -166,7 +166,7 @@ inductive DocInfo where
Turns an `Expr` into a pretty printed `CodeWithInfos`. Turns an `Expr` into a pretty printed `CodeWithInfos`.
-/ -/
def prettyPrintTerm (expr : Expr) : MetaM CodeWithInfos := do def prettyPrintTerm (expr : Expr) : MetaM CodeWithInfos := do
let ⟨fmt, infos⟩ ← PrettyPrinter.ppExprWithInfos expr let ⟨fmt, infos⟩ ← PrettyPrinter.ppExprWithInfos expr
let tt := TaggedText.prettyTagged fmt let tt := TaggedText.prettyTagged fmt
let ctx := { let ctx := {
env := ← getEnv env := ← getEnv
@ -177,9 +177,9 @@ def prettyPrintTerm (expr : Expr) : MetaM CodeWithInfos := do
fileMap := default, fileMap := default,
ngen := ← getNGen ngen := ← getNGen
} }
pure <| tagCodeInfos ctx infos tt return tagCodeInfos ctx infos tt
def isInstance (declName : Name) : MetaM Bool := do def isInstance (declName : Name) : MetaM Bool := do
pure <| (instanceExtension.getState (←getEnv)).instanceNames.contains declName return (instanceExtension.getState (← getEnv)).instanceNames.contains declName
end DocGen4.Process end DocGen4.Process

View File

@ -33,18 +33,18 @@ def valueToEq (v : DefinitionVal) : MetaM Expr := withLCtx {} {} do
let us := v.levelParams.map mkLevelParam let us := v.levelParams.map mkLevelParam
let type ← mkEq (mkAppN (Lean.mkConst v.name us) xs) body let type ← mkEq (mkAppN (Lean.mkConst v.name us) xs) body
let type ← mkForallFVars xs type let type ← mkForallFVars xs type
pure type return type
def DefinitionInfo.ofDefinitionVal (v : DefinitionVal) : MetaM DefinitionInfo := do def DefinitionInfo.ofDefinitionVal (v : DefinitionVal) : MetaM DefinitionInfo := do
let info ← Info.ofConstantVal v.toConstantVal let info ← Info.ofConstantVal v.toConstantVal
let isUnsafe := v.safety == DefinitionSafety.unsafe let isUnsafe := v.safety == DefinitionSafety.unsafe
let isNonComputable := isNoncomputable (←getEnv) v.name let isNonComputable := isNoncomputable (← getEnv) v.name
try try
let eqs? ← getEqnsFor? v.name let eqs? ← getEqnsFor? v.name
match eqs? with match eqs? with
| some eqs => | some eqs =>
let equations ← eqs.mapM processEq let equations ← eqs.mapM processEq
pure { return {
toInfo := info, toInfo := info,
isUnsafe, isUnsafe,
hints := v.hints, hints := v.hints,
@ -52,8 +52,8 @@ def DefinitionInfo.ofDefinitionVal (v : DefinitionVal) : MetaM DefinitionInfo :=
isNonComputable isNonComputable
} }
| none => | none =>
let equations := #[←prettyPrintTerm <| stripArgs (←valueToEq v)] let equations := #[← prettyPrintTerm <| stripArgs (← valueToEq v)]
pure { return {
toInfo := info, toInfo := info,
isUnsafe, isUnsafe,
hints := v.hints, hints := v.hints,
@ -61,8 +61,8 @@ def DefinitionInfo.ofDefinitionVal (v : DefinitionVal) : MetaM DefinitionInfo :=
isNonComputable isNonComputable
} }
catch err => catch err =>
IO.println s!"WARNING: Failed to calculate equational lemmata for {v.name}: {←err.toMessageData.toString}" IO.println s!"WARNING: Failed to calculate equational lemmata for {v.name}: {← err.toMessageData.toString}"
pure { return {
toInfo := info, toInfo := info,
isUnsafe, isUnsafe,
hints := v.hints, hints := v.hints,

View File

@ -99,7 +99,7 @@ def getDocString : DocInfo → Option String
| classInductiveInfo i => i.doc | classInductiveInfo i => i.doc
def isBlackListed (declName : Name) : MetaM Bool := do def isBlackListed (declName : Name) : MetaM Bool := do
match ←findDeclarationRanges? declName with match ← findDeclarationRanges? declName with
| some _ => | some _ =>
let env ← getEnv let env ← getEnv
pure (declName.isInternal) pure (declName.isInternal)
@ -108,7 +108,7 @@ def isBlackListed (declName : Name) : MetaM Bool := do
<||> isRec declName <||> isRec declName
<||> isMatcher declName <||> isMatcher declName
-- TODO: Evaluate whether filtering out declarations without range is sensible -- TODO: Evaluate whether filtering out declarations without range is sensible
| none => pure true | none => return true
-- TODO: Is this actually the best way? -- TODO: Is this actually the best way?
def isProjFn (declName : Name) : MetaM Bool := do def isProjFn (declName : Name) : MetaM Bool := do
@ -119,44 +119,44 @@ def isProjFn (declName : Name) : MetaM Bool := do
match getStructureInfo? env parent with match getStructureInfo? env parent with
| some i => | some i =>
match i.fieldNames.find? (· == name) with match i.fieldNames.find? (· == name) with
| some _ => pure true | some _ => return true
| none => pure false | none => return false
| none => panic! s!"{parent} is not a structure" | none => panic! s!"{parent} is not a structure"
else else
pure false return false
| _ => pure false | _ => return false
def ofConstant : (Name × ConstantInfo) → MetaM (Option DocInfo) := λ (name, info) => do def ofConstant : (Name × ConstantInfo) → MetaM (Option DocInfo) := fun (name, info) => do
if (←isBlackListed name) then if ← isBlackListed name then
return none return none
match info with match info with
| ConstantInfo.axiomInfo i => pure <| some <| axiomInfo (←AxiomInfo.ofAxiomVal i) | ConstantInfo.axiomInfo i => return some <| axiomInfo (← AxiomInfo.ofAxiomVal i)
| ConstantInfo.thmInfo i => pure <| some <| theoremInfo (←TheoremInfo.ofTheoremVal i) | ConstantInfo.thmInfo i => return some <| theoremInfo (← TheoremInfo.ofTheoremVal i)
| ConstantInfo.opaqueInfo i => pure <| some <| opaqueInfo (←OpaqueInfo.ofOpaqueVal i) | ConstantInfo.opaqueInfo i => return some <| opaqueInfo (← OpaqueInfo.ofOpaqueVal i)
| ConstantInfo.defnInfo i => | ConstantInfo.defnInfo i =>
if ←isProjFn i.name then if ← isProjFn i.name then
pure none return none
else else
if ←isInstance i.name then if ← isInstance i.name then
let info ← InstanceInfo.ofDefinitionVal i let info ← InstanceInfo.ofDefinitionVal i
pure <| some <| instanceInfo info return some <| instanceInfo info
else else
let info ← DefinitionInfo.ofDefinitionVal i let info ← DefinitionInfo.ofDefinitionVal i
pure <| some <| definitionInfo info return some <| definitionInfo info
| ConstantInfo.inductInfo i => | ConstantInfo.inductInfo i =>
let env ← getEnv let env ← getEnv
if isStructure env i.name then if isStructure env i.name then
if isClass env i.name then if isClass env i.name then
pure <| some <| classInfo (←ClassInfo.ofInductiveVal i) return some <| classInfo (← ClassInfo.ofInductiveVal i)
else else
pure <| some <| structureInfo (←StructureInfo.ofInductiveVal i) return some <| structureInfo (← StructureInfo.ofInductiveVal i)
else else
if isClass env i.name then if isClass env i.name then
pure <| some <| classInductiveInfo (←ClassInductiveInfo.ofInductiveVal i) return some <| classInductiveInfo (← ClassInductiveInfo.ofInductiveVal i)
else else
pure <| some <| inductiveInfo (←InductiveInfo.ofInductiveVal i) return some <| inductiveInfo (← InductiveInfo.ofInductiveVal i)
-- we ignore these for now -- we ignore these for now
| ConstantInfo.ctorInfo _ | ConstantInfo.recInfo _ | ConstantInfo.quotInfo _ => pure none | ConstantInfo.ctorInfo _ | ConstantInfo.recInfo _ | ConstantInfo.quotInfo _ => return none
def getKindDescription : DocInfo → String def getKindDescription : DocInfo → String
| axiomInfo i => if i.isUnsafe then "unsafe axiom" else "axiom" | axiomInfo i => if i.isUnsafe then "unsafe axiom" else "axiom"
@ -168,7 +168,7 @@ def getKindDescription : DocInfo → String
| DefinitionSafety.partial => "partial def" | DefinitionSafety.partial => "partial def"
| definitionInfo i => Id.run do | definitionInfo i => Id.run do
if i.hints.isAbbrev then if i.hints.isAbbrev then
pure "abbrev" return "abbrev"
else else
let mut modifiers := #[] let mut modifiers := #[]
if i.isUnsafe then if i.isUnsafe then
@ -177,7 +177,7 @@ def getKindDescription : DocInfo → String
modifiers := modifiers.push "noncomputable" modifiers := modifiers.push "noncomputable"
modifiers := modifiers.push "def" modifiers := modifiers.push "def"
pure <| String.intercalate " " modifiers.toList return String.intercalate " " modifiers.toList
| instanceInfo i => Id.run do | instanceInfo i => Id.run do
let mut modifiers := #[] let mut modifiers := #[]
if i.isUnsafe then if i.isUnsafe then
@ -186,7 +186,7 @@ def getKindDescription : DocInfo → String
modifiers := modifiers.push "noncomputable" modifiers := modifiers.push "noncomputable"
modifiers := modifiers.push "instance" modifiers := modifiers.push "instance"
pure <| String.intercalate " " modifiers.toList return String.intercalate " " modifiers.toList
| inductiveInfo i => if i.isUnsafe then "unsafe inductive" else "inductive" | inductiveInfo i => if i.isUnsafe then "unsafe inductive" else "inductive"
| structureInfo _ => "structure" | structureInfo _ => "structure"
| classInfo _ => "class" | classInfo _ => "class"

View File

@ -18,11 +18,11 @@ def getNLevels (name : Name) (levels: Nat) : Name :=
(components.drop (components.length - levels)).reverse.foldl (· ++ ·) Name.anonymous (components.drop (components.length - levels)).reverse.foldl (· ++ ·) Name.anonymous
inductive Hierarchy where inductive Hierarchy where
| node (name : Name) (isFile : Bool) (children : RBNode Name (λ _ => Hierarchy)) : Hierarchy | node (name : Name) (isFile : Bool) (children : RBNode Name (fun _ => Hierarchy)) : Hierarchy
instance : Inhabited Hierarchy := ⟨Hierarchy.node Name.anonymous false RBNode.leaf⟩ instance : Inhabited Hierarchy := ⟨Hierarchy.node Name.anonymous false RBNode.leaf⟩
abbrev HierarchyMap := RBNode Name (λ _ => Hierarchy) abbrev HierarchyMap := RBNode Name (fun _ => Hierarchy)
-- Everything in this namespace is adapted from stdlib's RBNode -- Everything in this namespace is adapted from stdlib's RBNode
namespace HierarchyMap namespace HierarchyMap
@ -100,23 +100,23 @@ def baseDirBlackList : HashSet String :=
partial def fromDirectoryAux (dir : System.FilePath) (previous : Name) : IO (Array Name) := do partial def fromDirectoryAux (dir : System.FilePath) (previous : Name) : IO (Array Name) := do
let mut children := #[] let mut children := #[]
for entry in ←System.FilePath.readDir dir do for entry in ← System.FilePath.readDir dir do
if (←entry.path.isDir) then if ← entry.path.isDir then
children := children ++ (←fromDirectoryAux entry.path (.str previous entry.fileName)) children := children ++ (← fromDirectoryAux entry.path (.str previous entry.fileName))
else if entry.path.extension = some "html" then else if entry.path.extension = some "html" then
children := children.push <| .str previous (entry.fileName.dropRight ".html".length) children := children.push <| .str previous (entry.fileName.dropRight ".html".length)
pure children return children
def fromDirectory (dir : System.FilePath) : IO Hierarchy := do def fromDirectory (dir : System.FilePath) : IO Hierarchy := do
let mut children := #[] let mut children := #[]
for entry in ←System.FilePath.readDir dir do for entry in ← System.FilePath.readDir dir do
if baseDirBlackList.contains entry.fileName then if baseDirBlackList.contains entry.fileName then
continue continue
else if ←entry.path.isDir then else if ← entry.path.isDir then
children := children ++ (←fromDirectoryAux entry.path (.mkSimple entry.fileName)) children := children ++ (← fromDirectoryAux entry.path (.mkSimple entry.fileName))
else if entry.path.extension = some "html" then else if entry.path.extension = some "html" then
children := children.push <| .mkSimple (entry.fileName.dropRight ".html".length) children := children.push <| .mkSimple (entry.fileName.dropRight ".html".length)
pure <| Hierarchy.fromArray children return Hierarchy.fromArray children
end Hierarchy end Hierarchy
end DocGen4 end DocGen4

View File

@ -20,8 +20,8 @@ def getConstructorType (ctor : Name) : MetaM Expr := do
def InductiveInfo.ofInductiveVal (v : InductiveVal) : MetaM InductiveInfo := do def InductiveInfo.ofInductiveVal (v : InductiveVal) : MetaM InductiveInfo := do
let info ← Info.ofConstantVal v.toConstantVal let info ← Info.ofConstantVal v.toConstantVal
let ctors ← v.ctors.mapM (λ name => do NameInfo.ofTypedName name (←getConstructorType name)) let ctors ← v.ctors.mapM (fun name => do NameInfo.ofTypedName name (← getConstructorType name))
pure { return {
toInfo := info, toInfo := info,
ctors, ctors,
isUnsafe := v.isUnsafe isUnsafe := v.isUnsafe

View File

@ -24,7 +24,7 @@ where
| .sort .zero => modify (·.push "_builtin_prop") | .sort .zero => modify (·.push "_builtin_prop")
| .sort (.succ _) => modify (·.push "_builtin_typeu") | .sort (.succ _) => modify (·.push "_builtin_typeu")
| .sort _ => modify (·.push "_builtin_sortu") | .sort _ => modify (·.push "_builtin_sortu")
| _ => pure () | _ => return ()
def InstanceInfo.ofDefinitionVal (v : DefinitionVal) : MetaM InstanceInfo := do def InstanceInfo.ofDefinitionVal (v : DefinitionVal) : MetaM InstanceInfo := do
let mut info ← DefinitionInfo.ofDefinitionVal v let mut info ← DefinitionInfo.ofDefinitionVal v
@ -32,7 +32,6 @@ def InstanceInfo.ofDefinitionVal (v : DefinitionVal) : MetaM InstanceInfo := do
if let some instAttr ← getDefaultInstance v.name className then if let some instAttr ← getDefaultInstance v.name className then
info := { info with attrs := info.attrs.push instAttr } info := { info with attrs := info.attrs.push instAttr }
let typeNames ← getInstanceTypes v.type let typeNames ← getInstanceTypes v.type
return { return {
toDefinitionInfo := info, toDefinitionInfo := info,
className, className,

View File

@ -13,15 +13,15 @@ open Lean Meta
def NameInfo.ofTypedName (n : Name) (t : Expr) : MetaM NameInfo := do def NameInfo.ofTypedName (n : Name) (t : Expr) : MetaM NameInfo := do
let env ← getEnv let env ← getEnv
pure { name := n, type := ←prettyPrintTerm t, doc := ←findDocString? env n} return { name := n, type := ← prettyPrintTerm t, doc := ← findDocString? env n}
partial def typeToArgsType (e : Expr) : (Array (Name × Expr × BinderInfo) × Expr) := partial def typeToArgsType (e : Expr) : (Array (Name × Expr × BinderInfo) × Expr) :=
let helper := λ name type body data => let helper := fun name type body data =>
-- Once we hit a name with a macro scope we stop traversing the expression -- Once we hit a name with a macro scope we stop traversing the expression
-- and print what is left after the : instead. The only exception -- and print what is left after the : instead. The only exception
-- to this is instances since these almost never have a name -- to this is instances since these almost never have a name
-- but should still be printed as arguments instead of after the :. -- but should still be printed as arguments instead of after the :.
if name.hasMacroScopes ∧ ¬data.isInstImplicit then if name.hasMacroScopes && !data.isInstImplicit then
(#[], e) (#[], e)
else else
let name := name.eraseMacroScopes let name := name.eraseMacroScopes
@ -35,15 +35,15 @@ partial def typeToArgsType (e : Expr) : (Array (Name × Expr × BinderInfo) × E
def Info.ofConstantVal (v : ConstantVal) : MetaM Info := do def Info.ofConstantVal (v : ConstantVal) : MetaM Info := do
let (args, type) := typeToArgsType v.type let (args, type) := typeToArgsType v.type
let args ← args.mapM (λ (n, e, b) => do pure <| Arg.mk n (←prettyPrintTerm e) b) let args ← args.mapM (fun (n, e, b) => do return Arg.mk n (← prettyPrintTerm e) b)
let nameInfo ← NameInfo.ofTypedName v.name type let nameInfo ← NameInfo.ofTypedName v.name type
match ←findDeclarationRanges? v.name with match ← findDeclarationRanges? v.name with
-- TODO: Maybe selection range is more relevant? Figure this out in the future -- TODO: Maybe selection range is more relevant? Figure this out in the future
| some range => pure { | some range => return {
toNameInfo := nameInfo, toNameInfo := nameInfo,
args, args,
declarationRange := range.range, declarationRange := range.range,
attrs := (←getAllAttributes v.name) attrs := ← getAllAttributes v.name
} }
| none => panic! s!"{v.name} is a declaration without position" | none => panic! s!"{v.name} is a declaration without position"

View File

@ -24,7 +24,7 @@ def OpaqueInfo.ofOpaqueVal (v : OpaqueVal) : MetaM OpaqueInfo := do
DefinitionSafety.unsafe DefinitionSafety.unsafe
else else
DefinitionSafety.safe DefinitionSafety.safe
pure { return {
toInfo := info, toInfo := info,
value, value,
definitionSafety definitionSafety

View File

@ -20,7 +20,7 @@ def dropArgs (type : Expr) (n : Nat) : (Expr × List (Name × Expr)) :=
let body := body.instantiate1 <| mkFVar ⟨name⟩ let body := body.instantiate1 <| mkFVar ⟨name⟩
let next := dropArgs body x let next := dropArgs body x
{ next with snd := (name, type) :: next.snd} { next with snd := (name, type) :: next.snd}
| _e, _x + 1 => panic! s!"No forallE left" | _, _ + 1 => panic! s!"No forallE left"
def getFieldTypes (struct : Name) (ctor : ConstructorVal) (parents : Nat) : MetaM (Array NameInfo) := do def getFieldTypes (struct : Name) (ctor : ConstructorVal) (parents : Nat) : MetaM (Array NameInfo) := do
let type := ctor.type let type := ctor.type
@ -28,8 +28,8 @@ def getFieldTypes (struct : Name) (ctor : ConstructorVal) (parents : Nat) : Meta
let (_, fields) := dropArgs fieldFunction (ctor.numFields - parents) let (_, fields) := dropArgs fieldFunction (ctor.numFields - parents)
let mut fieldInfos := #[] let mut fieldInfos := #[]
for (name, type) in fields do for (name, type) in fields do
fieldInfos := fieldInfos.push <| ←NameInfo.ofTypedName (struct.append name) type fieldInfos := fieldInfos.push <| ← NameInfo.ofTypedName (struct.append name) type
pure <| fieldInfos return fieldInfos
def StructureInfo.ofInductiveVal (v : InductiveVal) : MetaM StructureInfo := do def StructureInfo.ofInductiveVal (v : InductiveVal) : MetaM StructureInfo := do
let info ← Info.ofConstantVal v.toConstantVal let info ← Info.ofConstantVal v.toConstantVal
@ -40,14 +40,14 @@ def StructureInfo.ofInductiveVal (v : InductiveVal) : MetaM StructureInfo := do
match getStructureInfo? env v.name with match getStructureInfo? env v.name with
| some i => | some i =>
if i.fieldNames.size - parents.size > 0 then if i.fieldNames.size - parents.size > 0 then
pure { return {
toInfo := info, toInfo := info,
fieldInfo := (←getFieldTypes v.name ctorVal parents.size), fieldInfo := ← getFieldTypes v.name ctorVal parents.size,
parents, parents,
ctor ctor
} }
else else
pure { return {
toInfo := info, toInfo := info,
fieldInfo := #[], fieldInfo := #[],
parents, parents,

View File

@ -14,6 +14,6 @@ open Lean Meta
def TheoremInfo.ofTheoremVal (v : TheoremVal) : MetaM TheoremInfo := do def TheoremInfo.ofTheoremVal (v : TheoremVal) : MetaM TheoremInfo := do
let info ← Info.ofConstantVal v.toConstantVal let info ← Info.ofConstantVal v.toConstantVal
pure { toInfo := info } return { toInfo := info }
end DocGen4.Process end DocGen4.Process

View File

@ -8,7 +8,7 @@ def getTopLevelModules (p : Parsed) : IO (List String) := do
let topLevelModules := p.variableArgsAs! String |>.toList let topLevelModules := p.variableArgsAs! String |>.toList
if topLevelModules.length == 0 then if topLevelModules.length == 0 then
throw <| IO.userError "No topLevelModules provided." throw <| IO.userError "No topLevelModules provided."
pure topLevelModules return topLevelModules
def runSingleCmd (p : Parsed) : IO UInt32 := do def runSingleCmd (p : Parsed) : IO UInt32 := do
let relevantModules := [p.positionalArg! "module" |>.as! String |> String.toName] let relevantModules := [p.positionalArg! "module" |>.as! String |> String.toName]
@ -19,14 +19,14 @@ def runSingleCmd (p : Parsed) : IO UInt32 := do
IO.println "Outputting HTML" IO.println "Outputting HTML"
let baseConfig := getSimpleBaseContext hierarchy let baseConfig := getSimpleBaseContext hierarchy
htmlOutputResults baseConfig doc ws (p.hasFlag "ink") htmlOutputResults baseConfig doc ws (p.hasFlag "ink")
pure 0 return 0
| Except.error rc => pure rc | Except.error rc => pure rc
def runIndexCmd (_p : Parsed) : IO UInt32 := do def runIndexCmd (_p : Parsed) : IO UInt32 := do
let hierarchy ← Hierarchy.fromDirectory Output.basePath let hierarchy ← Hierarchy.fromDirectory Output.basePath
let baseConfig := getSimpleBaseContext hierarchy let baseConfig := getSimpleBaseContext hierarchy
htmlOutputIndex baseConfig htmlOutputIndex baseConfig
pure 0 return 0
def runGenCoreCmd (_p : Parsed) : IO UInt32 := do def runGenCoreCmd (_p : Parsed) : IO UInt32 := do
let res ← lakeSetup let res ← lakeSetup
@ -36,7 +36,7 @@ def runGenCoreCmd (_p : Parsed) : IO UInt32 := do
IO.println "Outputting HTML" IO.println "Outputting HTML"
let baseConfig := getSimpleBaseContext hierarchy let baseConfig := getSimpleBaseContext hierarchy
htmlOutputResults baseConfig doc ws (ink := False) htmlOutputResults baseConfig doc ws (ink := False)
pure 0 return 0
| Except.error rc => pure rc | Except.error rc => pure rc
def runDocGenCmd (_p : Parsed) : IO UInt32 := do def runDocGenCmd (_p : Parsed) : IO UInt32 := do